diff --git a/thys/Complx/ex/Examples.thy b/thys/Complx/ex/Examples.thy --- a/thys/Complx/ex/Examples.thy +++ b/thys/Complx/ex/Examples.thy @@ -1,267 +1,265 @@ (* * Copyright 2016, Data61, CSIRO * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(DATA61_BSD) *) section \Examples\ theory Examples imports "../OG_Syntax" begin record test = x :: nat y :: nat text \This is a sequence of two commands, the first being an assign protected by two guards. The guards use booleans as their faults.\ definition "test_guard \ \True\ (True, \\x=0\), (False, \(0::nat)=0\) \ \True\ \y := 0;; \True\ \x := 0" lemma "\, \ |\\<^bsub>/{True}\<^esub> COBEGIN test_guard \True\,\True\ \ \True\ \y:=0 \True\, \True\ COEND \True\,\True\" unfolding test_guard_def apply oghoare (*11 subgoals*) apply simp_all done definition "test_try_throw \ TRY \True\ \y := 0;; \True\ THROW CATCH \True\ \x := 0 END" subsection \Parameterized Examples\ subsubsection \Set Elements of an Array to Zero\ record Example1 = ex1_a :: "nat \ nat" lemma Example1: "\, \|\\<^bsub>/F\<^esub>\True\ COBEGIN SCHEME [0\iTrue\ \ex1_a:=\ex1_a (i:=0) \\ex1_a i=0\, \False\ COEND \\i < n. \ex1_a i = 0\, X" apply oghoare (* 7 subgoals *) apply (simp ; fail)+ done text \Same example but with a Call.\ definition "Example1'a \ \True\ \ex1_a:=\ex1_a (0:=0)" definition "Example1'b \ \True\ \ex1_a:=\ex1_a (1:=0)" definition "Example1' \ COBEGIN Example1'a \\ex1_a 0=0\, \False\ \ \True\ SCALL 0 0 \\ex1_a 1=0\, \False\ COEND" definition "\' = Map.empty(0 \ com Example1'b)" definition "\' = Map.empty(0 :: nat \ [ann Example1'b])" lemma Example1_proc_simp[unfolded Example1'b_def oghoare_simps]: "\' 0 = Some (com (Example1'b))" "\' 0 = Some ([ ann(Example1'b)])" "[ ann(Example1'b)]!0 = ann(Example1'b)" by (simp add: \'_def \'_def)+ lemma Example1': notes Example1_proc_simp[proc_simp] shows "\', \' |\\<^bsub>/F\<^esub> Example1' \\i < 2. \ex1_a i = 0\, \False\" unfolding Example1'_def apply simp unfolding Example1'a_def Example1'b_def apply oghoare (*12 subgoals*) apply simp+ using less_2_cases apply blast apply simp apply (erule disjE ; clarsimp) done type_synonym routine = nat text \Same example but with a Call.\ record Example2 = ex2_n :: "routine \ nat" ex2_a :: "nat \ string" definition Example2'n :: "routine \ (Example2, string \ nat, 'f) ann_com" where "Example2'n i \ \\ex2_n i= i\ \ex2_a:=\ex2_a((\ex2_n i):='''')" lemma Example2'n_map_of_simps[simp]: "i < n \ map_of (map (\i. ((p, i), g i)) [0..'' n \ map_of (map (\i. ((''f'', i), com (Example2'n i))) [0..'' n \ map_of (map (\i. ((''f'', i), [ann (Example2'n i)])) [0.. \'' n (''f'',i) = Some ( com(Example2'n i))" "i \'' n (''f'',i) = Some ([ ann(Example2'n i)])" "[ ann(Example2'n i)]!0 = ann(Example2'n i)" by (simp add: \''_def \''_def)+ lemmas Example2'n_proc_simp[proc_simp add] lemma Example2: notes Example2'n_proc_simp[proc_simp] shows "\'' n, \'' n |\\<^bsub>/F\<^esub>\True\ COBEGIN SCHEME [0\iTrue\ CALLX (\s. s\ex2_n:=(ex2_n s)(i:=i)\) \\ex2_n i = i\ (''f'', i) 0 (\s t. t\ex2_n:= (ex2_n t)(i:=(ex2_n s) i)\) (\x y. Skip) \\ex2_a (\ex2_n i)='''' \ \ex2_n i = i\ \\ex2_a i=''''\ \False\ \False\ \\ex2_a i=''''\, \False\ COEND \\i < n. \ex2_a i = ''''\, \False\" unfolding Example2'n_def ann_call_def call_def block_def apply oghoare (* 113 subgoals *) apply (clarsimp ; fail)+ done lemmas Example2'n_proc_simp[proc_simp del] text \Same example with lists as auxiliary variables.\ record Example2_list = ex2_A :: "nat list" lemma Example2_list: "\, \ |\\<^bsub>/F\<^esub>\n < length \ex2_A\ COBEGIN SCHEME [0\in < length \ex2_A\ \ex2_A:=\ex2_A[i:=0] \\ex2_A!i=0\,\False\ COEND \\i < n. \ex2_A!i = 0\, X" apply oghoare (*7 subgoals*) apply force+ done lemma exceptions_example: "\, \ |\\<^bsub>/F\<^esub> TRY \True \ \y := 0;; \ \y = 0 \ THROW CATCH \\y = 0\ \x := \y + 1 END \ \x = 1 \ \y = 0\, \False\" by oghoare simp_all lemma guard_example: "\, \ |\\<^bsub>/{42,66}\<^esub> \True\ (42, \\x=0\), (66, \\y=0\) \ \\x = 0\ \y := 0;; \True\ \x := 0 \ \x = 0\, \False\" apply oghoare (*6 subgoals*) apply simp_all done subsubsection \Peterson's mutex algorithm I (from Hoare-Parallel) \ text \Eike Best. "Semantics of Sequential and Parallel Programs", page 217.\ record Petersons_mutex_1 = pr1 :: nat pr2 :: nat in1 :: bool in2 :: bool hold :: nat lemma peterson_thread_1: "\, \ |\\<^bsub>/F\<^esub> \\pr1=0 \ \\in1\ WHILE True INV \\pr1=0 \ \\in1\ DO \\pr1=0 \ \\in1\ \\in1:=True,, \pr1:=1 \;; \\pr1=1 \ \in1\ \\hold:=1,, \pr1:=2 \;; \\pr1=2 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)\ AWAIT (\\in2 \ \(\hold=1)) THEN \pr1:=3 END;; \\pr1=3 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)\ \\in1:=False,,\pr1:=0\ OD \\pr1=0 \ \\in1\,\False\ " apply oghoare (*7 subgoals*) apply (((auto)[1]) ; fail)+ done lemma peterson_thread_2: "\, \ |\\<^bsub>/F\<^esub> \\pr2=0 \ \\in2\ WHILE True INV \\pr2=0 \ \\in2\ DO \\pr2=0 \ \\in2\ \\in2:=True,, \pr2:=1 \;; \\pr2=1 \ \in2\ \ \hold:=2,, \pr2:=2 \ ;; \\pr2=2 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))\ AWAIT (\\in1 \ \(\hold=2)) THEN \pr2:=3 END;; \\pr2=3 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))\ \\in2:=False,, \pr2:=0\ OD \\pr2=0 \ \\in2\,\False\ " apply oghoare (*7 subgoals*) - apply ((auto[1]) ; fail)+ - done + by auto lemma Petersons_mutex_1: "\, \ |\\<^bsub>/F\<^esub> \\pr1=0 \ \\in1 \ \pr2=0 \ \\in2 \ COBEGIN \\pr1=0 \ \\in1 \ WHILE True INV \\pr1=0 \ \\in1\ DO \\pr1=0 \ \\in1\ \ \in1:=True,, \pr1:=1 \;; \\pr1=1 \ \in1\ \ \hold:=1,, \pr1:=2 \;; \\pr1=2 \ \in1 \ (\hold=1 \ (\hold=2 \ \pr2=2))\ AWAIT (\\in2 \ \(\hold=1)) THEN \pr1:=3 END;; \\pr1=3 \ \in1 \ (\hold=1 \ (\hold=2 \ \pr2=2))\ \ \in1:=False,, \pr1:=0\ OD \\pr1=0 \ \\in1\,\False\ \ \\pr2=0 \ \\in2\ WHILE True INV \\pr2=0 \ \\in2\ DO \\pr2=0 \ \\in2\ \ \in2:=True,, \pr2:=1 \;; \\pr2=1 \ \in2\ \ \hold:=2,, \pr2:=2 \ ;; \\pr2=2 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))\ AWAIT (\\in1 \ \(\hold=2)) THEN \pr2:=3 END;; \\pr2=3 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))\ \ \in2:=False,, \pr2:=0\ OD \\pr2=0 \ \\in2\,\False\ COEND \\pr1=0 \ \\in1 \ \pr2=0 \ \\in2\,\False\" apply oghoare \ \81 verification conditions.\ - apply (((auto)[1]) ; fail)+ - done + by auto end diff --git a/thys/IP_Addresses/WordInterval.thy b/thys/IP_Addresses/WordInterval.thy --- a/thys/IP_Addresses/WordInterval.thy +++ b/thys/IP_Addresses/WordInterval.thy @@ -1,777 +1,777 @@ (* Title: WordInterval.thy Authors: Julius Michaelis, Cornelius Diekmann *) theory WordInterval imports Main "Word_Lib.Word_Lemmas" "Word_Lib.Next_and_Prev" begin section\WordInterval: Executable datatype for Machine Word Sets\ text\Stores ranges of machine words as interval. This has been proven quite efficient for IP Addresses.\ (*NOTE: All algorithms here use a straight-forward implementation. There is a lot of room for improving the computation complexity, for example by making the WordInterval a balanced, sorted tree.*) subsection\Syntax\ context notes [[typedef_overloaded]] begin datatype ('a::len) wordinterval = WordInterval "('a::len) word" \ \start (inclusive)\ "('a::len) word" \ \end (inclusive)\ | RangeUnion "'a wordinterval" "'a wordinterval" end subsection\Semantics\ fun wordinterval_to_set :: "'a::len wordinterval \ ('a::len word) set" where "wordinterval_to_set (WordInterval start end) = {start .. end}" | "wordinterval_to_set (RangeUnion r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" (*Note: The runtime of all the operations could be improved, for example by keeping the tree sorted and balanced.*) subsection\Basic operations\ text\\\\\ fun wordinterval_element :: "'a::len word \ 'a::len wordinterval \ bool" where "wordinterval_element el (WordInterval s e) \ s \ el \ el \ e" | "wordinterval_element el (RangeUnion r1 r2) \ wordinterval_element el r1 \ wordinterval_element el r2" lemma wordinterval_element_set_eq[simp]: "wordinterval_element el rg = (el \ wordinterval_to_set rg)" by(induction rg rule: wordinterval_element.induct) simp_all definition wordinterval_union :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_union r1 r2 = RangeUnion r1 r2" lemma wordinterval_union_set_eq[simp]: "wordinterval_to_set (wordinterval_union r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" unfolding wordinterval_union_def by simp fun wordinterval_empty :: "'a::len wordinterval \ bool" where "wordinterval_empty (WordInterval s e) \ e < s" | "wordinterval_empty (RangeUnion r1 r2) \ wordinterval_empty r1 \ wordinterval_empty r2" lemma wordinterval_empty_set_eq[simp]: "wordinterval_empty r \ wordinterval_to_set r = {}" by(induction r) auto definition Empty_WordInterval :: "'a::len wordinterval" where "Empty_WordInterval \ WordInterval 1 0" lemma wordinterval_empty_Empty_WordInterval: "wordinterval_empty Empty_WordInterval" by(simp add: Empty_WordInterval_def) lemma Empty_WordInterval_set_eq[simp]: "wordinterval_to_set Empty_WordInterval = {}" by(simp add: Empty_WordInterval_def) subsection\WordInterval and Lists\ text\A list of \(start, end)\ tuples.\ text\wordinterval to list\ fun wi2l :: "'a::len wordinterval \ ('a::len word \ 'a::len word) list" where "wi2l (RangeUnion r1 r2) = wi2l r1 @ wi2l r2" | "wi2l (WordInterval s e) = (if e < s then [] else [(s,e)])" text\list to wordinterval\ fun l2wi :: "('a::len word \ 'a word) list \ 'a wordinterval" where "l2wi [] = Empty_WordInterval" | "l2wi [(s,e)] = (WordInterval s e)" | "l2wi ((s,e)#rs) = (RangeUnion (WordInterval s e) (l2wi rs))" lemma l2wi_append: "wordinterval_to_set (l2wi (l1@l2)) = wordinterval_to_set (l2wi l1) \ wordinterval_to_set (l2wi l2)" proof(induction l1 arbitrary: l2 rule:l2wi.induct) case 1 thus ?case by simp next case (2 s e l2) thus ?case by (cases l2) simp_all next case 3 thus ?case by force qed lemma l2wi_wi2l[simp]: "wordinterval_to_set (l2wi (wi2l r)) = wordinterval_to_set r" by(induction r) (simp_all add: l2wi_append) lemma l2wi: "wordinterval_to_set (l2wi l) = (\ (i,j) \ set l. {i .. j})" by(induction l rule: l2wi.induct, simp_all) lemma wi2l: "(\(i,j)\set (wi2l r). {i .. j}) = wordinterval_to_set r" by(induction r rule: wi2l.induct, simp_all) lemma l2wi_remdups[simp]: "wordinterval_to_set (l2wi (remdups ls)) = wordinterval_to_set (l2wi ls)" by(simp add: l2wi) lemma wi2l_empty[simp]: "wi2l Empty_WordInterval = []" unfolding Empty_WordInterval_def by simp subsection\Optimizing and minimizing @{typ "('a::len) wordinterval"}s\ text\Removing empty intervals\ context begin fun wordinterval_optimize_empty :: "'a::len wordinterval \ 'a wordinterval" where "wordinterval_optimize_empty (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty r1o then r2o else if wordinterval_empty r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty r = r" lemma wordinterval_optimize_empty_set_eq[simp]: "wordinterval_to_set (wordinterval_optimize_empty r) = wordinterval_to_set r" by(induction r) (simp_all add: Let_def) lemma wordinterval_optimize_empty_double: "wordinterval_optimize_empty (wordinterval_optimize_empty r) = wordinterval_optimize_empty r" by(induction r) (simp_all add: Let_def) private fun wordinterval_empty_shallow :: "'a::len wordinterval \ bool" where "wordinterval_empty_shallow (WordInterval s e) \ e < s" | "wordinterval_empty_shallow (RangeUnion _ _) \ False" private lemma helper_optimize_shallow: "wordinterval_empty_shallow (wordinterval_optimize_empty r) = wordinterval_empty (wordinterval_optimize_empty r)" by(induction r) fastforce+ private fun wordinterval_optimize_empty2 where "wordinterval_optimize_empty2 (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty_shallow r1o then r2o else if wordinterval_empty_shallow r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty2 r = r" lemma wordinterval_optimize_empty_code[code_unfold]: "wordinterval_optimize_empty = wordinterval_optimize_empty2" by (subst fun_eq_iff, clarify, rename_tac r, induct_tac r) (unfold wordinterval_optimize_empty.simps wordinterval_optimize_empty2.simps Let_def helper_optimize_shallow, simp_all) end text\Merging overlapping intervals\ context begin private definition disjoint :: "'a set \ 'a set \ bool" where "disjoint A B \ A \ B = {}" private primrec interval_of :: "('a::len) word \ 'a word \ 'a word set" where "interval_of (s,e) = {s .. e}" declare interval_of.simps[simp del] private definition disjoint_intervals :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "disjoint_intervals A B \ disjoint (interval_of A) (interval_of B)" private definition not_disjoint_intervals :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "not_disjoint_intervals A B \ \ disjoint (interval_of A) (interval_of B)" private lemma [code]: "not_disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s \ e' \ s' \ e \ s \ e \ s' \ e')" apply(cases A, cases B) apply(simp add: not_disjoint_intervals_def interval_of.simps disjoint_def) done private lemma [code]: "disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s > e' \ s' > e \ s > e \ s' > e')" apply(cases A, cases B) apply(simp add: disjoint_intervals_def interval_of.simps disjoint_def) by fastforce text\BEGIN merging overlapping intervals\ (*result has no empty intervals and all are disjoint. merging things such as [1,7] [8,10] would still be possible*) private fun merge_overlap :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_overlap s [] = [s]" | "merge_overlap (s,e) ((s',e')#ss) = ( if not_disjoint_intervals (s,e) (s',e') then (min s s', max e e')#ss else (s',e')#merge_overlap (s,e) ss)" private lemma not_disjoint_union: fixes s :: "('a::len) word" shows "\ disjoint {s..e} {s'..e'} \ {s..e} \ {s'..e'} = {min s s' .. max e e'}" by(auto simp add: disjoint_def min_def max_def) private lemma disjoint_subset: "disjoint A B \ A \ B \ C \ A \ C" unfolding disjoint_def by blast private lemma merge_overlap_helper1: "interval_of A \ (\s \ set ss. interval_of s) \ (\s \ set (merge_overlap A ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(drule_tac C="(\x\set xs. interval_of x)" in disjoint_subset) apply(simp_all) done private lemma merge_overlap_helper2: "\s'\set ss. \ disjoint (interval_of A) (interval_of s') \ interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_overlap A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(simp) by blast private lemma merge_overlap_length: "\s' \ set ss. \ disjoint (interval_of A) (interval_of s') \ length (merge_overlap A ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) done lemma "merge_overlap (1:: 16 word,2) [(1, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(2, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(3, 7)] = [(3, 7), (1,2)]" by eval private function listwordinterval_compress :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_compress [] = []" | "listwordinterval_compress (s#ss) = ( if \s' \ set ss. disjoint_intervals s s' then s#listwordinterval_compress ss else listwordinterval_compress (merge_overlap s ss))" by(pat_completeness, auto) private termination listwordinterval_compress apply (relation "measure length") apply(rule wf_measure) apply(simp) using disjoint_intervals_def merge_overlap_length by fastforce private lemma listwordinterval_compress: "(\s \ set (listwordinterval_compress ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_compress.induct) apply(simp) apply(simp) apply(intro impI) apply(simp add: disjoint_intervals_def) apply(drule merge_overlap_helper2) apply(simp) done lemma "listwordinterval_compress [(1::32 word,3), (8,10), (2,5), (3,7)] = [(8, 10), (1, 7)]" by eval private lemma A_in_listwordinterval_compress: "A \ set (listwordinterval_compress ss) \ interval_of A \ (\s \ set ss. interval_of s)" using listwordinterval_compress by blast private lemma listwordinterval_compress_disjoint: "A \ set (listwordinterval_compress ss) \ B \ set (listwordinterval_compress ss) \ A \ B \ disjoint (interval_of A) (interval_of B)" apply(induction ss arbitrary: rule: listwordinterval_compress.induct) apply(simp) apply(simp split: if_split_asm) apply(elim disjE) apply(simp_all) apply(simp_all add: disjoint_intervals_def disjoint_def) apply(blast dest: A_in_listwordinterval_compress)+ done text\END merging overlapping intervals\ text\BEGIN merging adjacent intervals\ private fun merge_adjacent :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_adjacent s [] = [s]" | "merge_adjacent (s,e) ((s',e')#ss) = ( if s \e \ s' \ e' \ word_next e = s' then (s, e')#ss else if s \e \ s' \ e' \ word_next e' = s then (s', e)#ss else (s',e')#merge_adjacent (s,e) ss)" private lemma merge_adjacent_helper: "interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_adjacent A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: interval_of.simps) apply(intro impI conjI) apply (metis Un_assoc word_adjacent_union) apply(elim conjE) apply(drule(2) word_adjacent_union) subgoal by (blast) subgoal by (metis word_adjacent_union Un_assoc) by blast private lemma merge_adjacent_length: "\(s', e')\set ss. s \ e \ s' \ e' \ (word_next e = s' \ word_next e' = s) \ length (merge_adjacent (s,e) ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(case_tac x) - applysimp + apply simp by blast private function listwordinterval_adjacent :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_adjacent [] = []" | "listwordinterval_adjacent ((s,e)#ss) = ( if \(s',e') \ set ss. \ (s \e \ s' \ e' \ (word_next e = s' \ word_next e' = s)) then (s,e)#listwordinterval_adjacent ss else listwordinterval_adjacent (merge_adjacent (s,e) ss))" by(pat_completeness, auto) private termination listwordinterval_adjacent apply (relation "measure length") apply(rule wf_measure) apply(simp) apply(simp) using merge_adjacent_length by fastforce private lemma listwordinterval_adjacent: "(\s \ set (listwordinterval_adjacent ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_adjacent.induct) apply(simp) apply(simp add: merge_adjacent_helper) done lemma "listwordinterval_adjacent [(1::16 word, 3), (5, 10), (10,10), (4,4)] = [(10, 10), (1, 10)]" by eval text\END merging adjacent intervals\ definition wordinterval_compress :: "('a::len) wordinterval \ 'a wordinterval" where "wordinterval_compress r \ l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))" text\Correctness: Compression preserves semantics\ lemma wordinterval_compress: "wordinterval_to_set (wordinterval_compress r) = wordinterval_to_set r" unfolding wordinterval_compress_def proof - have interval_of': "interval_of s = (case s of (s,e) \ {s .. e})" for s by (cases s) (simp add: interval_of.simps) have "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = (\x\set (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))). interval_of x)" by (force simp: interval_of' l2wi) also have "\ = (\s\set (wi2l (wordinterval_optimize_empty r)). interval_of s)" by(simp add: listwordinterval_compress listwordinterval_adjacent) also have "\ = (\(i, j)\set (wi2l (wordinterval_optimize_empty r)). {i..j})" by(simp add: interval_of') also have "\ = wordinterval_to_set r" by(simp add: wi2l) finally show "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = wordinterval_to_set r" . qed end text\Example\ lemma "(wi2l \ (wordinterval_compress :: 32 wordinterval \ 32 wordinterval) \ l2wi) [(70, 80001), (0,0), (150, 8000), (1,3), (42,41), (3,7), (56, 200), (8,10)] = [(56, 80001), (0, 10)]" by eval lemma "wordinterval_compress (RangeUnion (RangeUnion (WordInterval (1::32 word) 5) (WordInterval 8 10)) (WordInterval 3 7)) = WordInterval 1 10" by eval subsection\Further operations\ text\\\\\ definition wordinterval_Union :: "('a::len) wordinterval list \ 'a wordinterval" where "wordinterval_Union ws = wordinterval_compress (foldr wordinterval_union ws Empty_WordInterval)" lemma wordinterval_Union: "wordinterval_to_set (wordinterval_Union ws) = (\ w \ (set ws). wordinterval_to_set w)" by(induction ws) (simp_all add: wordinterval_compress wordinterval_Union_def) context begin private fun wordinterval_setminus' :: "'a::len wordinterval \ 'a wordinterval \ 'a wordinterval" where "wordinterval_setminus' (WordInterval s e) (WordInterval ms me) = ( if s > e \ ms > me then WordInterval s e else if me \ e then WordInterval (if ms = 0 then 1 else s) (min e (word_prev ms)) else if ms \ s then WordInterval (max s (word_next me)) (if me = - 1 then 0 else e) else RangeUnion (WordInterval (if ms = 0 then 1 else s) (word_prev ms)) (WordInterval (word_next me) (if me = - 1 then 0 else e)) )" | "wordinterval_setminus' (RangeUnion r1 r2) t = RangeUnion (wordinterval_setminus' r1 t) (wordinterval_setminus' r2 t)"| "wordinterval_setminus' t (RangeUnion r1 r2) = wordinterval_setminus' (wordinterval_setminus' t r1) r2" private lemma wordinterval_setminus'_rr_set_eq: "wordinterval_to_set(wordinterval_setminus' (WordInterval s e) (WordInterval ms me)) = wordinterval_to_set (WordInterval s e) - wordinterval_to_set (WordInterval ms me)" apply(simp only: wordinterval_setminus'.simps) apply(case_tac "e < s") apply simp apply(case_tac "me < ms") apply simp apply(case_tac [!] "e \ me") apply(case_tac [!] "ms = 0") apply(case_tac [!] "ms \ s") apply(case_tac [!] "me = - 1") apply(simp_all add: word_next_unfold word_prev_unfold min_def max_def) apply(safe) apply(auto) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) done private lemma wordinterval_setminus'_set_eq: "wordinterval_to_set (wordinterval_setminus' r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" apply(induction rule: wordinterval_setminus'.induct) using wordinterval_setminus'_rr_set_eq apply blast apply auto done lemma wordinterval_setminus'_empty_struct: "wordinterval_empty r2 \ wordinterval_setminus' r1 r2 = r1" by(induction r1 r2 rule: wordinterval_setminus'.induct) auto definition wordinterval_setminus :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_setminus r1 r2 = wordinterval_compress (wordinterval_setminus' r1 r2)" lemma wordinterval_setminus_set_eq[simp]: "wordinterval_to_set (wordinterval_setminus r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" by(simp add: wordinterval_setminus_def wordinterval_compress wordinterval_setminus'_set_eq) end definition wordinterval_UNIV :: "'a::len wordinterval" where "wordinterval_UNIV \ WordInterval 0 (- 1)" lemma wordinterval_UNIV_set_eq[simp]: "wordinterval_to_set wordinterval_UNIV = UNIV" unfolding wordinterval_UNIV_def using max_word_max by fastforce fun wordinterval_invert :: "'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_invert r = wordinterval_setminus wordinterval_UNIV r" lemma wordinterval_invert_set_eq[simp]: "wordinterval_to_set (wordinterval_invert r) = UNIV - wordinterval_to_set r" by(auto) lemma wordinterval_invert_UNIV_empty: "wordinterval_empty (wordinterval_invert wordinterval_UNIV)" by simp lemma wi2l_univ[simp]: "wi2l wordinterval_UNIV = [(0, - 1)]" unfolding wordinterval_UNIV_def by simp text\\\\\ context begin private lemma "{(s::nat) .. e} \ {s' .. e'} = {} \ s > e' \ s' > e \ s > e \ s' > e'" by simp linarith private fun wordinterval_intersection' :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection' (WordInterval s e) (WordInterval s' e') = ( if s > e \ s' > e' \ s > e' \ s' > e \ s > e \ s' > e' then Empty_WordInterval else WordInterval (max s s') (min e e') )" | "wordinterval_intersection' (RangeUnion r1 r2) t = RangeUnion (wordinterval_intersection' r1 t) (wordinterval_intersection' r2 t)"| "wordinterval_intersection' t (RangeUnion r1 r2) = RangeUnion (wordinterval_intersection' t r1) (wordinterval_intersection' t r2)" private lemma wordinterval_intersection'_set_eq: "wordinterval_to_set (wordinterval_intersection' r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(induction r1 r2 rule: wordinterval_intersection'.induct) (auto) lemma "wordinterval_intersection' (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = RangeUnion (RangeUnion (WordInterval 1 3) (WordInterval 1 0)) (WordInterval 1 3)" by eval definition wordinterval_intersection :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection r1 r2 \ wordinterval_compress (wordinterval_intersection' r1 r2)" lemma wordinterval_intersection_set_eq[simp]: "wordinterval_to_set (wordinterval_intersection r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(simp add: wordinterval_intersection_def wordinterval_compress wordinterval_intersection'_set_eq) lemma "wordinterval_intersection (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = WordInterval 1 3" by eval end definition wordinterval_subset :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_subset r1 r2 \ wordinterval_empty (wordinterval_setminus r1 r2)" lemma wordinterval_subset_set_eq[simp]: "wordinterval_subset r1 r2 = (wordinterval_to_set r1 \ wordinterval_to_set r2)" unfolding wordinterval_subset_def by simp definition wordinterval_eq :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_eq r1 r2 = (wordinterval_subset r1 r2 \ wordinterval_subset r2 r1)" lemma wordinterval_eq_set_eq: "wordinterval_eq r1 r2 \ wordinterval_to_set r1 = wordinterval_to_set r2" unfolding wordinterval_eq_def by auto thm iffD1[OF wordinterval_eq_set_eq] (*declare iffD1[OF wordinterval_eq_set_eq, simp]*) lemma wordinterval_eq_comm: "wordinterval_eq r1 r2 \ wordinterval_eq r2 r1" unfolding wordinterval_eq_def by fast lemma wordinterval_to_set_alt: "wordinterval_to_set r = {x. wordinterval_element x r}" unfolding wordinterval_element_set_eq by blast lemma wordinterval_un_empty: "wordinterval_empty r1 \ wordinterval_eq (wordinterval_union r1 r2) r2" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_un_emty_b: "wordinterval_empty r2 \ wordinterval_eq (wordinterval_union r1 r2) r1" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_Diff_triv: "wordinterval_empty (wordinterval_intersection a b) \ wordinterval_eq (wordinterval_setminus a b) a" unfolding wordinterval_eq_set_eq by simp blast text\A size of the datatype, does not correspond to the cardinality of the corresponding set\ fun wordinterval_size :: "('a::len) wordinterval \ nat" where "wordinterval_size (RangeUnion a b) = wordinterval_size a + wordinterval_size b" | "wordinterval_size (WordInterval s e) = (if s \ e then 1 else 0)" lemma wordinterval_size_length: "wordinterval_size r = length (wi2l r)" by(induction r) (auto) lemma Ex_wordinterval_nonempty: "\x::('a::len wordinterval). y \ wordinterval_to_set x" proof show "y \ wordinterval_to_set wordinterval_UNIV" by simp qed lemma wordinterval_eq_reflp: "reflp wordinterval_eq" apply(rule reflpI) by(simp only: wordinterval_eq_set_eq) lemma wordintervalt_eq_symp: "symp wordinterval_eq" apply(rule sympI) by(simp add: wordinterval_eq_comm) lemma wordinterval_eq_transp: "transp wordinterval_eq" apply(rule transpI) by(simp only: wordinterval_eq_set_eq) lemma wordinterval_eq_equivp: "equivp wordinterval_eq" by (auto intro: equivpI wordinterval_eq_reflp wordintervalt_eq_symp wordinterval_eq_transp) text\The smallest element in the interval\ definition is_lowest_element :: "'a::ord \ 'a set \ bool" where "is_lowest_element x S = (x \ S \ (\y\S. y \ x \ y = x))" lemma fixes x :: "'a :: complete_lattice" assumes "x \ S" shows " x = Inf S \ is_lowest_element x S" using assms apply(simp add: is_lowest_element_def) by (simp add: Inf_lower eq_iff) lemma fixes x :: "'a :: linorder" assumes "finite S" and "x \ S" shows "is_lowest_element x S \ x = Min S" apply(rule) subgoal apply(simp add: is_lowest_element_def) apply(subst Min_eqI[symmetric]) using assms by(auto) by (metis Min.coboundedI assms(1) assms(2) dual_order.antisym is_lowest_element_def) text\Smallest element in the interval\ fun wordinterval_lowest_element :: "'a::len wordinterval \ 'a word option" where "wordinterval_lowest_element (WordInterval s e) = (if s \ e then Some s else None)" | "wordinterval_lowest_element (RangeUnion A B) = (case (wordinterval_lowest_element A, wordinterval_lowest_element B) of (Some a, Some b) \ Some (if a < b then a else b) | (None, Some b) \ Some b | (Some a, None) \ Some a | (None, None) \ None)" lemma wordinterval_lowest_none_empty: "wordinterval_lowest_element r = None \ wordinterval_empty r" proof(induction r) case WordInterval thus ?case by simp next case RangeUnion thus ?case by fastforce qed lemma wordinterval_lowest_element_correct_A: "wordinterval_lowest_element r = Some x \ is_lowest_element x (wordinterval_to_set r)" unfolding is_lowest_element_def apply(induction r arbitrary: x rule: wordinterval_lowest_element.induct) apply(rename_tac rs re x, case_tac "rs \ re", auto)[1] apply(subst(asm) wordinterval_lowest_element.simps(2)) apply(rename_tac A B x) apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(simp_all add: wordinterval_lowest_none_empty)[3] apply fastforce done lemma wordinterval_lowest_element_set_eq: assumes "\ wordinterval_empty r" shows "(wordinterval_lowest_element r = Some x) = (is_lowest_element x (wordinterval_to_set r))" (*unfolding is_lowest_element_def*) proof(rule iffI) assume "wordinterval_lowest_element r = Some x" thus "is_lowest_element x (wordinterval_to_set r)" using wordinterval_lowest_element_correct_A wordinterval_lowest_none_empty by simp next assume "is_lowest_element x (wordinterval_to_set r)" with assms show "(wordinterval_lowest_element r = Some x)" proof(induction r arbitrary: x rule: wordinterval_lowest_element.induct) case 1 thus ?case by(simp add: is_lowest_element_def) next case (2 A B x) have is_lowest_RangeUnion: "is_lowest_element x (wordinterval_to_set A \ wordinterval_to_set B) \ is_lowest_element x (wordinterval_to_set A) \ is_lowest_element x (wordinterval_to_set B)" by(simp add: is_lowest_element_def) (*why \ A B?*) have wordinterval_lowest_element_RangeUnion: "\a b A B. wordinterval_lowest_element A = Some a \ wordinterval_lowest_element B = Some b \ wordinterval_lowest_element (RangeUnion A B) = Some (min a b)" by(auto dest!: wordinterval_lowest_element_correct_A simp add: is_lowest_element_def min_def) from 2 show ?case apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(auto simp add: is_lowest_element_def)[3] apply(subgoal_tac "\ wordinterval_empty A \ \ wordinterval_empty B") prefer 2 using arg_cong[where f = Not, OF wordinterval_lowest_none_empty] apply blast apply(drule(1) wordinterval_lowest_element_RangeUnion) apply(simp split: option.split_asm add: min_def) apply(drule is_lowest_RangeUnion) apply(elim disjE) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) using wordinterval_lowest_element_correct_A[simplified is_lowest_element_def] by (metis Un_iff not_le) qed qed text\Cardinality approximation for @{typ "('a::len) wordinterval"}s\ context begin lemma card_atLeastAtMost_word: fixes s::"('a::len) word" shows "card {s..e} = Suc (unat e) - (unat s)" apply(cases "s > e") apply(simp) apply(subst(asm) Word.word_less_nat_alt) apply simp apply(subst upto_enum_set_conv2[symmetric]) apply(subst List.card_set) apply(simp add: remdups_enum_upto) done fun wordinterval_card :: "('a::len) wordinterval \ nat" where "wordinterval_card (WordInterval s e) = Suc (unat e) - (unat s)" | "wordinterval_card (RangeUnion a b) = wordinterval_card a + wordinterval_card b" lemma wordinterval_card: "wordinterval_card r \ card (wordinterval_to_set r)" proof(induction r) case WordInterval thus ?case by (simp add: card_atLeastAtMost_word) next case (RangeUnion r1 r2) have "card (wordinterval_to_set r1 \ wordinterval_to_set r2) \ card (wordinterval_to_set r1) + card (wordinterval_to_set r2)" using Finite_Set.card_Un_le by blast with RangeUnion show ?case by(simp) qed text\With @{thm wordinterval_compress} it should be possible to get the exact cardinality\ end end diff --git a/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy b/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy --- a/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy +++ b/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy @@ -1,203 +1,203 @@ theory IpAddresses imports IP_Addresses.IP_Address_toString IP_Addresses.CIDR_Split "../Common/WordInterval_Lists" begin \ \Misc\ (*we dont't have an empty ip space, but a space which only contains the 0 address. We will use the option type to denote the empty space in some functions.*) lemma "ipset_from_cidr (ipv4addr_of_dotdecimal (0, 0, 0, 0)) 33 = {0}" by(simp add: ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def ipset_from_cidr_large_pfxlen) (*helper we use for spoofing protection specification*) definition all_but_those_ips :: "('i::len word \ nat) list \ ('i word \ nat) list" where "all_but_those_ips cidrips = cidr_split (wordinterval_invert (l2wi (map ipcidr_to_interval cidrips)))" lemma all_but_those_ips: "ipcidr_union_set (set (all_but_those_ips cidrips)) = UNIV - (\ (ip,n) \ set cidrips. ipset_from_cidr ip n)" - applysimp + apply simp unfolding ipcidr_union_set_uncurry all_but_those_ips_def apply(simp add: cidr_split_prefix) apply(simp add: l2wi) apply(simp add: ipcidr_to_interval_def) using ipset_from_cidr_ipcidr_to_interval by blast section\IPv4 Addresses\ subsection\IPv4 Addresses in IPTables Notation (how we parse it)\ context notes [[typedef_overloaded]] begin datatype 'i ipt_iprange = \ \Singleton IP Address\ IpAddr "'i::len word" \ \CIDR notation: addr/xx\ | IpAddrNetmask "'i word" nat \ \-m iprange --src-range a.b.c.d-e.f.g.h\ | IpAddrRange "'i word" "'i word" (*the range is inclusive*) end fun ipt_iprange_to_set :: "'i::len ipt_iprange \ 'i word set" where "ipt_iprange_to_set (IpAddrNetmask base m) = ipset_from_cidr base m" | "ipt_iprange_to_set (IpAddr ip) = { ip }" | "ipt_iprange_to_set (IpAddrRange ip1 ip2) = { ip1 .. ip2 }" text\@{term ipt_iprange_to_set} can only represent an empty set if it is an empty range.\ lemma ipt_iprange_to_set_nonempty: "ipt_iprange_to_set ip = {} \ (\ip1 ip2. ip = IpAddrRange ip1 ip2 \ ip1 > ip2)" apply(cases ip) apply(simp; fail) apply(simp add: ipset_from_cidr_alt bitmagic_zeroLast_leq_or1Last; fail) apply(simp add:linorder_not_le; fail) done text\maybe this is necessary as code equation?\ context includes bit_operations_syntax begin lemma element_ipt_iprange_to_set[code_unfold]: "(addr::'i::len word) \ ipt_iprange_to_set X = ( case X of (IpAddrNetmask pre len) \ (pre AND ((mask len) << (len_of (TYPE('i)) - len))) \ addr \ addr \ pre OR (mask (len_of (TYPE('i)) - len)) | IpAddr ip \ (addr = ip) | IpAddrRange ip1 ip2 \ ip1 \ addr \ ip2 \ addr)" apply(cases X) apply(simp; fail) apply(simp add: ipset_from_cidr_alt; fail) apply(simp; fail) done end lemma ipt_iprange_to_set_uncurry_IpAddrNetmask: "ipt_iprange_to_set (uncurry IpAddrNetmask a) = uncurry ipset_from_cidr a" by(simp split: uncurry_splits) text\IP address ranges to \(start, end)\ notation\ fun ipt_iprange_to_interval :: "'i::len ipt_iprange \ ('i word \ 'i word)" where "ipt_iprange_to_interval (IpAddr addr) = (addr, addr)" | "ipt_iprange_to_interval (IpAddrNetmask pre len) = ipcidr_to_interval (pre, len)" | "ipt_iprange_to_interval (IpAddrRange ip1 ip2) = (ip1, ip2)" lemma ipt_iprange_to_interval: "ipt_iprange_to_interval ip = (s,e) \ {s .. e} = ipt_iprange_to_set ip" apply(cases ip) apply(auto simp add: ipcidr_to_interval) done text\A list of IP address ranges to a @{typ "'i::len wordinterval"}. The nice thing is: the usual set operations are defined on this type. We can use the existing function @{const l2wi_intersect} if we want the intersection of the supplied list\ lemma "wordinterval_to_set (l2wi_intersect (map ipt_iprange_to_interval ips)) = (\ ip \ set ips. ipt_iprange_to_set ip)" apply(simp add: l2wi_intersect) using ipt_iprange_to_interval by blast text\We can use @{const l2wi} if we want the union of the supplied list\ lemma "wordinterval_to_set (l2wi (map ipt_iprange_to_interval ips)) = (\ ip \ set ips. ipt_iprange_to_set ip)" apply(simp add: l2wi) using ipt_iprange_to_interval by blast text\A list of (negated) IP address to a @{typ "'i::len wordinterval"}.\ definition ipt_iprange_negation_type_to_br_intersect :: "'i::len ipt_iprange negation_type list \ 'i wordinterval" where "ipt_iprange_negation_type_to_br_intersect l = l2wi_negation_type_intersect (NegPos_map ipt_iprange_to_interval l)" lemma ipt_iprange_negation_type_to_br_intersect: "wordinterval_to_set (ipt_iprange_negation_type_to_br_intersect l) = (\ ip \ set (getPos l). ipt_iprange_to_set ip) - (\ ip \ set (getNeg l). ipt_iprange_to_set ip)" apply(simp add: ipt_iprange_negation_type_to_br_intersect_def l2wi_negation_type_intersect NegPos_map_simps) using ipt_iprange_to_interval by blast text\The @{typ "'i::len wordinterval"} can be translated back into a list of IP ranges. If a list of intervals is enough, we can use @{const wi2l}. If we need it in @{typ "'i::len ipt_iprange"}, we can use this function.\ definition wi_2_cidr_ipt_iprange_list :: "'i::len wordinterval \ 'i ipt_iprange list" where "wi_2_cidr_ipt_iprange_list r = map (uncurry IpAddrNetmask) (cidr_split r)" lemma wi_2_cidr_ipt_iprange_list: "(\ ip \ set (wi_2_cidr_ipt_iprange_list r). ipt_iprange_to_set ip) = wordinterval_to_set r" proof - have "(\ ip \ set (wi_2_cidr_ipt_iprange_list r). ipt_iprange_to_set ip) = (\x\set (cidr_split r). uncurry ipset_from_cidr x)" unfolding wi_2_cidr_ipt_iprange_list_def by force thus ?thesis using cidr_split_prefix by metis qed text\For example, this allows the following transformation\ definition ipt_iprange_compress :: "'i::len ipt_iprange negation_type list \ 'i ipt_iprange list" where "ipt_iprange_compress = wi_2_cidr_ipt_iprange_list \ ipt_iprange_negation_type_to_br_intersect" lemma ipt_iprange_compress: "(\ ip \ set (ipt_iprange_compress l). ipt_iprange_to_set ip) = (\ ip \ set (getPos l). ipt_iprange_to_set ip) - (\ ip \ set (getNeg l). ipt_iprange_to_set ip)" by (metis wi_2_cidr_ipt_iprange_list comp_apply ipt_iprange_compress_def ipt_iprange_negation_type_to_br_intersect) definition normalized_cidr_ip :: "'i::len ipt_iprange \ bool" where "normalized_cidr_ip ip \ case ip of IpAddrNetmask _ _ \ True | _ \ False" lemma wi_2_cidr_ipt_iprange_list_normalized_IpAddrNetmask: "\a'\set (wi_2_cidr_ipt_iprange_list as). normalized_cidr_ip a'" apply(clarify) apply(simp add: wi_2_cidr_ipt_iprange_list_def normalized_cidr_ip_def) by force lemma ipt_iprange_compress_normalized_IpAddrNetmask: "\a'\set (ipt_iprange_compress as). normalized_cidr_ip a'" by(simp add: ipt_iprange_compress_def wi_2_cidr_ipt_iprange_list_normalized_IpAddrNetmask) definition ipt_iprange_to_cidr :: "'i::len ipt_iprange \ ('i word \ nat) list" where "ipt_iprange_to_cidr ips = cidr_split (iprange_interval (ipt_iprange_to_interval ips))" lemma ipt_ipvange_to_cidr: "ipcidr_union_set (set (ipt_iprange_to_cidr ips)) = (ipt_iprange_to_set ips)" apply(simp add: ipt_iprange_to_cidr_def) apply(simp add: ipcidr_union_set_uncurry) apply(case_tac "(ipt_iprange_to_interval ips)") apply(simp add: ipt_iprange_to_interval cidr_split_prefix_single) done (* actually, these are toString pretty printing helpers*) definition interval_to_wi_to_ipt_iprange :: "'i::len word \ 'i word \ 'i ipt_iprange" where "interval_to_wi_to_ipt_iprange s e \ if s = e then IpAddr s else case cidr_split (WordInterval s e) of [(ip,nmask)] \ IpAddrNetmask ip nmask | _ \ IpAddrRange s e" lemma interval_to_wi_to_ipt_ipv4range: "ipt_iprange_to_set (interval_to_wi_to_ipt_iprange s e) = {s..e}" proof - from cidr_split_prefix_single[of s e] have "cidr_split (WordInterval s e) = [(a, b)] \ ipset_from_cidr a b = {s..e}" for a b by(simp add: iprange_interval.simps) thus ?thesis by(simp add: interval_to_wi_to_ipt_iprange_def split: list.split) qed fun wi_to_ipt_iprange :: "'i::len wordinterval \ 'i ipt_iprange list" where "wi_to_ipt_iprange (WordInterval s e) = (if s > e then [] else [interval_to_wi_to_ipt_iprange s e])" | "wi_to_ipt_iprange (RangeUnion a b) = wi_to_ipt_iprange a @ wi_to_ipt_iprange b" lemma wi_to_ipt_ipv4range: "\(set (map ipt_iprange_to_set (wi_to_ipt_iprange wi))) = wordinterval_to_set wi" apply(induction wi) apply(simp add: interval_to_wi_to_ipt_ipv4range) apply(simp) done end diff --git a/thys/List_Update/BIT.thy b/thys/List_Update/BIT.thy --- a/thys/List_Update/BIT.thy +++ b/thys/List_Update/BIT.thy @@ -1,2158 +1,2158 @@ (* Title: Competitive Analysis of BIT Author: Max Haslbeck *) section "BIT: an Online Algorithm for the List Update Problem" theory BIT imports Bit_Strings MTF2_Effects begin abbreviation "config'' A qs init n == config_rand A init (take n qs)" lemma sum_my: fixes f g::"'b \ 'a::ab_group_add" assumes "finite A" "finite B" shows "(\x\A. f x) - (\x\B. g x) = (\x\(A \ B). f x - g x) + (\x\A-B. f x) - (\x\B-A. g x)" proof - have "finite (A-B)" and "finite (A\B)" and "finite (B-A)" and "finite (B\A)" using assms by auto note finites=this have "(A-B) \ ( (A\B) ) = {}" and "(B-A) \ ( (B\A) ) = {}" by auto note inters=this have commute: "A\B=B\A" by auto have "A = (A-B) \ (A\B)" and "B = (B-A) \ ( (B\A))" by auto then have "(\x\A. f x) - (\x\B. g x) = (\x\(A-B) \ (A\B). f x) - (\x\(B-A) \ (B\A). g x)" by auto also have "\ = ( (\x\(A-B). f x) + (\x\(A\B). f x) - (\x\(A-B)\(A\B). f x) ) -( (\x\(B-A). g x) + (\x\(B\A). g x) - (\x\(B-A)\(B\A). g x))" using sum_Un[where ?f="f",OF finites(1) finites(2)] sum_Un[where ?f="g",OF finites(3) finites(4)] by(simp) also have "\ = ( (\x\(A-B). f x) + (\x\(A\B). f x) ) - (\x\(B-A). g x) - (\x\(B\A). g x) " using inters by auto also have "\ = (\x\(A-B). f x) - (\x\(A\B). g x) + (\x\(A\B). f x) - (\x\(B-A). g x) " using commute by auto also have "\ = (\x\(A\B). f x - g x) +(\x\(A-B). f x) - (\x\(B-A). g x)" using sum_subtractf[of f g "(A\B)"] by auto finally show ?thesis . qed lemma sum_my2: "(\x\A. f x = g x) \ (\x\A. f x) = (\x\A. g x)" by auto subsection "Definition of BIT" definition BIT_init :: "('a state,bool list * 'a list)alg_on_init" where "BIT_init init = map_pmf (\l. (l,init)) (bv (length init))" lemma "~ deterministic_init BIT_init" unfolding deterministic_init_def BIT_init_def apply(auto) apply(intro exI[where x="[a]"]) \ \comment in a proof\ by(auto simp: UNIV_bool set_pmf_bernoulli) definition BIT_step :: "('a state, bool list * 'a list, 'a, answer)alg_on_step" where "BIT_step s q = ( let a=((if (fst (snd s))!(index (snd (snd s)) q) then 0 else (length (fst s))),[]) in return_pmf (a , (flip (index (snd (snd s)) q) (fst (snd s)), snd (snd s))))" lemma "deterministic_step BIT_step" unfolding deterministic_step_def BIT_step_def by simp abbreviation BIT :: "('a state, bool list*'a list, 'a, answer)alg_on_rand" where "BIT == (BIT_init, BIT_step)" subsection "Properties of BIT's state distribution" lemma BIT_no_paid: "\((free,paid),_) \ (BIT_step s q). paid=[]" unfolding BIT_step_def by(auto) subsubsection "About the Internal State" term "(config'_rand (BIT_init, BIT_step) s0 qs) " lemma config'_n_init: fixes qs init n shows "map_pmf (snd \ snd) (config'_rand (BIT_init, BIT_step) init qs) = map_pmf (snd \ snd) init" apply (induct qs arbitrary: init) by (simp_all add: map_pmf_def bind_assoc_pmf BIT_step_def bind_return_pmf ) lemma config_n_init: "map_pmf (snd \ snd) (config_rand (BIT_init, BIT_step) s0 qs) = return_pmf s0" using config'_n_init[of "((fst (BIT_init, BIT_step) s0) \ (\is. return_pmf (s0, is)))"] by (simp_all add: map_pmf_def bind_assoc_pmf bind_return_pmf BIT_init_def ) lemma config_n_init2: "\(_,(_,x)) \ set_pmf (config_rand (BIT_init, BIT_step) init qs). x = init" proof (rule, goal_cases) case (1 z) then have 1: "snd(snd z) \ (snd \ snd) ` set_pmf (config_rand (BIT_init, BIT_step) init qs)" by force have "(snd \ snd) ` set_pmf (config_rand (BIT_init, BIT_step) init qs) = set_pmf (map_pmf (snd \ snd) (config_rand (BIT_init, BIT_step) init qs))" by(simp) also have "\ = {init}" apply(simp only: config_n_init) by simp finally have "snd(snd z) = init" using 1 by auto then show ?case by auto qed lemma config_n_init3: "\x \ set_pmf (config_rand (BIT_init, BIT_step) init qs). snd (snd x) = init" using config_n_init2 by(simp add: split_def) lemma config'_n_bv: fixes qs init n shows " map_pmf (snd \ snd) init = return_pmf s0 \ map_pmf (fst \ snd) init = bv (length s0) \ map_pmf (snd \ snd) (config'_rand (BIT_init, BIT_step) init qs) = return_pmf s0 \ map_pmf (fst \ snd) (config'_rand (BIT_init, BIT_step) init qs) = bv (length s0)" proof (induct qs arbitrary: init) case (Cons r rs) from Cons(2) have a: "map_pmf (snd \ snd) (init \ (\s. snd (BIT_init, BIT_step) s r \ (\(a, is'). return_pmf (step (fst s) r a, is')))) = return_pmf s0" apply(simp add: BIT_step_def) by (simp_all add: map_pmf_def bind_assoc_pmf BIT_step_def bind_return_pmf ) then have b: "\z\set_pmf (init \ (\s. snd (BIT_init, BIT_step) s r \ (\(a, is'). return_pmf (step (fst s) r a, is')))). snd (snd z) = s0" by (metis (mono_tags, lifting) comp_eq_dest_lhs map_pmf_eq_return_pmf_iff) show ?case apply(simp only: config'_rand.simps) proof (rule Cons(1), goal_cases) case 2 have "map_pmf (fst \ snd) (init \ (\s. snd (BIT_init, BIT_step) s r \ (\(a, is'). return_pmf (step (fst s) r a, is')))) = map_pmf (flip (index s0 r)) (bv (length s0))" using b apply(simp add: BIT_step_def Cons(3)[symmetric] bind_return_pmf map_pmf_def bind_assoc_pmf ) apply(rule bind_pmf_cong) apply(simp) by(simp add: inv_flip_bv) also have "\ = bv (length s0)" using inv_flip_bv by auto finally show ?case . qed (fact) qed simp lemma config_n_bv_2: "map_pmf (snd \ snd) (config_rand (BIT_init, BIT_step) s0 qs) = return_pmf s0 \ map_pmf (fst \ snd) (config_rand (BIT_init, BIT_step) s0 qs) = bv (length s0)" apply(rule config'_n_bv) by(simp_all add: bind_return_pmf map_pmf_def bind_assoc_pmf bind_return_pmf' BIT_init_def) lemma config_n_bv: "map_pmf (fst \ snd) (config_rand (BIT_init, BIT_step) s0 qs) = bv (length s0)" using config_n_bv_2 by auto lemma config_n_fst_init_length: "\(_,(x,_)) \ set_pmf (config_rand (BIT_init, BIT_step) s0 qs). length x = length s0" proof fix x::"('a list \ (bool list \ 'a list))" assume ass:"x \ set_pmf (config_rand (BIT_init, BIT_step) s0 qs)" let ?a="fst (snd x)" from ass have "(fst x,(?a,snd (snd x))) \ set_pmf (config_rand (BIT_init, BIT_step) s0 qs)" by auto with ass have "?a \ (fst \ snd) ` set_pmf (config_rand (BIT_init, BIT_step) s0 qs)" by force then have "?a \ set_pmf (map_pmf (fst \ snd) (config_rand (BIT_init, BIT_step) s0 qs))" by auto then have "?a \ bv (length s0)" by(simp only: config_n_bv) then have "length ?a = length s0" by (auto simp: len_bv_n) then show "case x of (uu_, xa, uua_) \ length xa = length s0" by(simp add: split_def) qed lemma config_n_fst_init_length2: "\x \ set_pmf (config_rand (BIT_init, BIT_step) s0 qs). length (fst (snd x)) = length s0" using config_n_fst_init_length by(simp add: split_def) lemma fperms: "finite {x::'a list. length x = length init \ distinct x \ set x = set init}" apply(rule finite_subset[where B="{xs. set xs \ set init \ length xs \ length init}"]) apply(force) apply(rule finite_lists_length_le) by auto lemma finite_config_BIT: assumes [simp]: "distinct init" shows "finite (set_pmf (config_rand (BIT_init, BIT_step) init qs))" (is "finite ?D") proof - have a: "(fst \ snd) ` ?D \ {x. length x = length init}" using config_n_fst_init_length2 by force have c: "(snd \ snd) ` ?D = {init}" proof - have "(snd \ snd) ` set_pmf (config_rand (BIT_init, BIT_step) init qs) = set_pmf (map_pmf (snd \ snd) (config_rand (BIT_init, BIT_step) init qs))" by(simp) also have "\ = {init}" apply(subst config_n_init) by simp finally show ?thesis . qed from a c have d: "snd ` ?D \ {x. length x = length init} \ {init}" by force have b: "fst ` ?D \ {x. length x = length init \ distinct x \ set x = set init}" using config_rand by fastforce from b d have "?D \ {x. length x = length init \ distinct x \ set x = set init} \ ({x. length x = length init} \ {init})" by auto then show ?thesis apply (rule finite_subset) apply(rule finite_cartesian_product) apply(rule fperms) apply(rule finite_cartesian_product) apply (rule bitstrings_finite) by(simp) qed subsection "BIT is $1.75$-competitive (a combinatorial proof)" subsubsection "Definition of the Locale and Helper Functions" locale BIT_Off = fixes acts :: "answer list" fixes qs :: "'a list" fixes init :: "'a list" assumes dist_init[simp]: "distinct init" assumes len_acts: "length acts = length qs" begin lemma setinit: "(index init) ` set init = {0.. {x. (a \ x)}) = set as" by fastforce have 2: "(\a. Suc (index as a)) ` set as = (\a. Suc a) ` ((index as) ` set as )" by auto show ?case apply(simp add: 1 2 iH) by auto qed simp definition free_A :: "nat list" where (* free exchanges of A *) "free_A = map fst acts" definition paid_A' :: "nat list list" where (* paid exchanges of A' *) "paid_A' = map snd acts" definition paid_A :: "nat list list" where (* paid exchanges of A *) "paid_A = map (filter (\x. Suc x < length init)) paid_A'" lemma len_paid_A[simp]: "length paid_A = length qs" unfolding paid_A_def paid_A'_def using len_acts by auto lemma len_paid_A'[simp]: "length paid_A' = length qs" unfolding paid_A'_def using len_acts by auto lemma paidAnm_inbound: "n < length paid_A \ m < length(paid_A!n) \ (Suc ((paid_A!n)!(length (paid_A ! n) - Suc m))) < length init" proof - assume "n < length paid_A" then have "n < length paid_A'" by auto then have a: "(paid_A!n) = filter (\x. Suc x < length init) (paid_A' ! n)" unfolding paid_A_def by auto let ?filtered="(filter (\x. Suc x < length init) (paid_A' ! n))" assume mtt: "m < length (paid_A!n)" with a have "(length (paid_A ! n) - Suc m) < length ?filtered" by auto with nth_mem have b: "Suc(?filtered ! (length (paid_A ! n) - Suc m)) < length init" by force show "Suc (paid_A ! n ! (length (paid_A ! n) - Suc m)) < length init" using a b by auto qed fun s_A' :: "nat \ 'a list" where "s_A' 0 = init" | "s_A'(Suc n) = step (s_A' n) (qs!n) (free_A!n, paid_A'!n)" lemma length_s_A'[simp]: "length(s_A' n) = length init" by (induction n) simp_all lemma dist_s_A'[simp]: "distinct(s_A' n)" by(induction n) (simp_all add: step_def) lemma set_s_A'[simp]: "set(s_A' n) = set init" by(induction n) (simp_all add: step_def) fun s_A :: "nat \ 'a list" where "s_A 0 = init" | "s_A(Suc n) = step (s_A n) (qs!n) (free_A!n, paid_A!n)" lemma length_s_A[simp]: "length(s_A n) = length init" by (induction n) simp_all lemma dist_s_A[simp]: "distinct(s_A n)" by(induction n) (simp_all add: step_def) lemma set_s_A[simp]: "set(s_A n) = set init" by(induction n) (simp_all add: step_def) lemma cost_paidAA': "n < length paid_A' \ length (paid_A!n) \ length (paid_A'!n)" unfolding paid_A_def by simp lemma swaps_filtered: "swaps (filter (\x. Suc x < length xs) ys) xs = swaps (ys) xs" apply (induct ys) by auto lemma sAsA': "n < length paid_A' \ s_A' n = s_A n" proof (induct n) case (Suc m) have " s_A' (Suc m) = mtf2 (free_A!m) (qs!m) (swaps (paid_A'!m) (s_A' m))" by (simp add: step_def) also from Suc(2) have "\ = mtf2 (free_A!m) (qs!m) (swaps (paid_A!m) (s_A' m))" unfolding paid_A_def by (simp only: nth_map swaps_filtered[where xs="s_A' m", simplified]) also have "\ = mtf2 (free_A!m) (qs!m) (swaps (paid_A!m) (s_A m))" using Suc by auto also have "\ = s_A (Suc m)" by (simp add: step_def) finally show ?case . qed simp lemma sAsA'': "n < length qs \ s_A n = s_A' n" using sAsA' by auto definition t_BIT :: "nat \ real" where (* BIT's cost in nth step *) "t_BIT n = T_on_rand_n BIT init qs n" definition T_BIT :: "nat \ real" where (* BIT's cost in first n steps *) "T_BIT n = (\i int" where "c_A n = index (swaps (paid_A!n) (s_A n)) (qs!n) + 1" definition f_A :: "nat \ int" where "f_A n = min (free_A!n) (index (swaps (paid_A!n) (s_A n)) (qs!n))" definition p_A :: "nat \ int" where "p_A n = size(paid_A!n)" definition t_A :: "nat \ int" where "t_A n = c_A n + p_A n" definition c_A' :: "nat \ int" where "c_A' n = index (swaps (paid_A'!n) (s_A' n)) (qs!n) + 1" definition p_A' :: "nat \ int" where "p_A' n = size(paid_A'!n)" definition t_A' :: "nat \ int" where "t_A' n = c_A' n + p_A' n" lemma t_A_A'_leq: "n < length paid_A' \ t_A n \ t_A' n" unfolding t_A_def t_A'_def c_A_def c_A'_def p_A_def p_A'_def apply(simp add: sAsA') unfolding paid_A_def by (simp add: swaps_filtered[where xs="(s_A n)", simplified]) definition T_A' :: "nat \ int" where "T_A' n = (\i int" where "T_A n = (\i length paid_A' \ T_A n \ T_A' n" unfolding T_A'_def T_A_def apply(rule sum_mono) by (simp add: t_A_A'_leq) lemma T_A_A'_leq': "n \ length qs \ T_A n \ T_A' n" using T_A_A'_leq by auto fun s'_A :: "nat \ nat \ 'a list" where "s'_A n 0 = s_A n" | "(s'_A n (Suc m)) = swap ((paid_A ! n)!(length (paid_A ! n) -(Suc m)) ) (s'_A n m)" lemma set_s'_A[simp]: "set (s'_A n m) = set init" apply(induct m) by(auto) lemma len_s'_A[simp]: "length (s'_A n m) = length init" apply(induct m) by(auto) lemma distperm_s'_A[simp]: "dist_perm (s'_A n m) init" apply(induct m) by auto lemma s'A_m_le: "m \ (length (paid_A ! n)) \ swaps (drop (length (paid_A ! n) - m) (paid_A ! n)) (s_A n) = s'_A n m" apply(induct m) apply(simp) proof - fix m assume iH: "(m \ length (paid_A ! n) \ swaps (drop (length (paid_A ! n) - m) (paid_A ! n)) (s_A n) = s'_A n m)" assume Suc: "Suc m \ length (paid_A ! n)" then have "m \ length (paid_A ! n)" by auto with iH have x: "swaps (drop (length (paid_A ! n) - m) (paid_A ! n)) (s_A n) = s'_A n m" by auto from Suc have mlen: "(length (paid_A ! n) - Suc m) < length (paid_A ! n)" by auto let ?l="length (paid_A ! n) - Suc m" let ?Sucl="length (paid_A ! n) - m" have Sucl: "Suc ?l = ?Sucl" using Suc by auto from mlen have yu: "((paid_A ! n)! ?l ) # (drop (Suc ?l) (paid_A ! n)) = (drop ?l (paid_A ! n))" by (rule Cons_nth_drop_Suc) from Suc have "s'_A n (Suc m) = swap ((paid_A ! n)!(length (paid_A ! n) - (Suc m)) ) (s'_A n m)" by auto also have "\ = swap ((paid_A ! n)!(length (paid_A ! n) - (Suc m)) ) (swaps (drop (length (paid_A ! n) - m) (paid_A ! n)) (s_A n))" by(simp only: x) also have "\ = (swaps (((paid_A ! n)!(length (paid_A ! n) - (Suc m)) ) # (drop (length (paid_A ! n) - m) (paid_A ! n))) (s_A n))" by auto also have "\ = (swaps (((paid_A ! n)! ?l ) # (drop (Suc ?l) (paid_A ! n))) (s_A n))" using Sucl by auto also from mlen have "\ = (swaps ((drop ?l (paid_A ! n))) (s_A n))" by (simp only: yu) finally have " s'_A n (Suc m) = swaps (drop (length (paid_A ! n) - Suc m) (paid_A ! n)) (s_A n)" . then show " swaps (drop (length (paid_A ! n) - Suc m) (paid_A ! n)) (s_A n) = s'_A n (Suc m)" by auto qed lemma s'A_m: "swaps (paid_A ! n) (s_A n) = s'_A n (length (paid_A ! n))" using s'A_m_le[of "(length (paid_A ! n))" "n", simplified] by auto definition gebub :: "nat \ nat \ nat" where "gebub n m = index init ((s'_A n m)!(Suc ((paid_A!n)!(length (paid_A ! n) - Suc m))))" lemma gebub_inBound: assumes 1: " n < length paid_A " and 2: "m < length (paid_A ! n)" shows "gebub n m < length init" proof - have "Suc (paid_A ! n ! (length (paid_A ! n) - Suc m)) < length (s'_A n m)" using paidAnm_inbound[OF 1 2] by auto then have "s'_A n m ! Suc (paid_A ! n ! (length (paid_A ! n) - Suc m)) \ set (s'_A n m)" by (rule nth_mem) then show ?thesis unfolding gebub_def using setinit by auto qed subsubsection "The Potential Function" fun phi :: "nat \'a list\ (bool list \ 'a list) \ real" ("\") where "phi n (c,(b,_)) = (\(x,y)\(Inv c (s_A n)). (if b!(index init y) then 2 else 1))" lemma phi': "phi n z = (\(x,y)\(Inv (fst z) (s_A n)). (if (fst (snd z))!(index init y) then 2 else 1))" proof - have "phi n z = phi n (fst z, (fst(snd z),snd(snd z)))" by (metis prod.collapse) also have "\ = (\(x,y)\(Inv (fst z) (s_A n)). (if (fst (snd z))!(index init y) then 2 else 1))" by(simp del: prod.collapse) finally show ?thesis . qed lemma Inv_empty2: "length d = 0 \ Inv c d = {}" unfolding Inv_def before_in_def by(auto) corollary Inv_empty3: "length init = 0 \ Inv c (s_A n) = {}" apply(rule Inv_empty2) by (metis length_s_A) lemma phi_empty2: "length init = 0 \ phi n (c,(b,i)) = 0" apply(simp only: phi.simps Inv_empty3) by auto lemma phi_nonzero: "phi n (c,(b,i)) \ 0" by (simp add: sum_nonneg split_def) (* definition of the potential function! *) definition Phi :: "nat \ real" ("\") where "Phi n = E( map_pmf (\ n) (config'' BIT qs init n))" definition PhiPlus :: "nat \ real" ("\\<^sup>+") where "PhiPlus n = (let nextconfig = bind_pmf (config'' BIT qs init n) (\(s,is). bind_pmf (BIT_step (s,is) (qs!n)) (\(a,nis). return_pmf (step s (qs!n) a,nis)) ) in E( map_pmf (phi (Suc n)) nextconfig) )" lemma PhiPlus_is_Phi_Suc: "n PhiPlus n = Phi (Suc n)" unfolding PhiPlus_def Phi_def apply (simp add: bind_return_pmf map_pmf_def bind_assoc_pmf split_def take_Suc_conv_app_nth ) apply(simp add: config'_rand_snoc) by(simp add: bind_assoc_pmf split_def bind_return_pmf) lemma phi0: "Phi 0 = 0" unfolding Phi_def by (simp add: bind_return_pmf map_pmf_def bind_assoc_pmf BIT_init_def) lemma phi_pos: "Phi n \ 0" unfolding Phi_def apply(rule E_nonneg_fun) using phi_nonzero by auto subsubsection "Helper lemmas" lemma swap_subs: "dist_perm X Y \ Inv X (swap z Y) \ Inv X Y \ {(Y ! z, Y ! Suc z)}" proof - assume "dist_perm X Y" note aj = Inv_swap[OF this, of z] show "Inv X (swap z Y) \ Inv X Y \ {(Y ! z, Y ! Suc z)}" proof cases assume c1: "Suc z < length X" show "Inv X (swap z Y) \ Inv X Y \ {(Y ! z, Y ! Suc z)}" proof cases assume "Y ! z < Y ! Suc z in X" with c1 have "Inv X (swap z Y) = Inv X Y \ {(Y ! z, Y ! Suc z)}" using aj by auto then show "Inv X (swap z Y) \ Inv X Y \ {(Y ! z, Y ! Suc z)}" by auto next assume "~ Y ! z < Y ! Suc z in X" with c1 have "Inv X (swap z Y) = Inv X Y - {(Y ! Suc z, Y ! z)}" using aj by auto then show "Inv X (swap z Y) \ Inv X Y \ {(Y ! z, Y ! Suc z)}" by auto qed next assume "~ (Suc z < length X)" then have "Inv X (swap z Y) = Inv X Y" using aj by auto then show "Inv X (swap z Y) \ Inv X Y \ {(Y ! z, Y ! Suc z)}" by auto qed qed subsubsection "InvOf" term "Inv" (* BIT A *) abbreviation "InvOf y bits as \ {(x,y)|x. x < y in bits \ y < x in as}" lemma "InvOf y xs ys = {(x,y)|x. (x,y)\Inv xs ys}" unfolding Inv_def by auto lemma "InvOf y xs ys \ Inv xs ys" unfolding Inv_def by auto lemma numberofIsbeschr: assumes distxsys: "dist_perm xs ys" and yinxs: "y \ set xs" shows "index xs y \ index ys y + card (InvOf y xs ys)" (is "?iBit \ ?iA + card ?I") proof - from assms have distinctxs: "distinct xs" and distinctys: "distinct ys" and yinys: "y \ set ys" by auto let ?A="fst ` ?I" have aha: "card ?A = card ?I" apply(rule card_image) unfolding inj_on_def by(auto) have "?A \ (before y xs)" by(auto) have "?A \ (after y ys)" by auto have "finite (before y ys)" by auto have bef: "(before y xs) - ?A \ before y ys" apply(auto) proof - fix x assume a: "x < y in xs" assume " x \ fst ` {(x, y) |x. x < y in xs \ y < x in ys}" then have "~ (x < y in xs \ y < x in ys)" by force with a have d: "~ y < x in ys" by auto from a have "x \ set xs" by (rule before_in_setD1) with distxsys have b: "x \ set ys" by auto from a have "y \ set xs" by (rule before_in_setD2) with distxsys have c: "y \ set ys" by auto from a have e: "~ x = y" unfolding before_in_def by auto have "(\ y < x in ys) = (x < y in ys \ y = x)" apply(rule not_before_in) using b c by auto with d e show "x < y in ys" by auto qed have "(index xs y) - card (InvOf y xs ys) = card (before y xs) - card ?A" by(simp only: aha card_before[OF distinctxs yinxs]) also have "\ = card ((before y xs) - ?A)" apply(rule card_Diff_subset[symmetric]) by auto also have "\ \ card (before y ys)" apply(rule card_mono) apply(simp) apply(rule bef) done also have "\ = (index ys y)" by(simp only: card_before[OF distinctys yinys]) finally have "index xs y - card ?I \ index ys y" . then show "index xs y \ index ys y + card ?I" by auto qed lemma "length init = 0 \ length xs = length init \ t xs q (mf, sws) = 1 + length sws" unfolding t_def by(auto) lemma integr_index: "integrable (measure_pmf (config'' (BIT_init, BIT_step) qs init n)) (\(s, is). real (Suc (index s (qs ! n))))" apply(rule measure_pmf.integrable_const_bound[where B="Suc (length init)"]) apply(simp add: split_def) apply (metis (mono_tags) index_le_size AE_measure_pmf_iff config_rand_length) by (auto) subsubsection "Upper Bound on the Cost of BIT" lemma t_BIT_ub2: "(qs!n) \ set init \ t_BIT n \ Suc(size init)" apply(simp add: t_BIT_def t_def BIT_step_def) apply(simp add: bind_return_pmf) proof (goal_cases) case 1 note qs=this let ?D = "(config'' (BIT_init, BIT_step) qs init n)" have absch: "(\x\ set_pmf ?D. ((\(s,is). real (Suc (index s (qs ! n)))) x) \ ((\(is,s). Suc (length init)) x))" proof (rule ballI, goal_cases) case (1 x) from 1 config_rand_length have f1: "length (fst x) = length init" by fastforce from 1 config_rand_set have 2: "set (fst x) = set init" by fastforce from qs 2 have "(qs!n) \ set (fst x)" by auto then show ?case using f1 by (simp add: split_def) qed have "integrable (measure_pmf (config'' (BIT_init, BIT_step) qs init n)) (\(s, is). Suc (length init))" by(simp) have "E(bind_pmf ?D (\(s, is). return_pmf (real (Suc (index s (qs ! n)))))) = E(map_pmf (\(s, is). real (Suc (index s (qs ! n)))) ?D)" by(simp add: split_def map_pmf_def) also have "\ \ E(map_pmf (\(s, is). Suc (length init)) ?D)" apply (rule E_mono3) apply(fact integr_index) apply(simp) using absch by auto also have "\ = Suc (length init)" by(simp add: split_def) finally show ?case by(simp add: map_pmf_def bind_assoc_pmf bind_return_pmf split_def) qed lemma t_BIT_ub: "(qs!n) \ set init \ t_BIT n \ size init" apply(simp add: t_BIT_def t_def BIT_step_def) apply(simp add: bind_return_pmf) proof (goal_cases) case 1 note qs=this let ?D = "(config'' (BIT_init, BIT_step) qs init n)" have absch: "(\x\ set_pmf ?D. ((\(s, is). real (Suc (index s (qs ! n)))) x) \ ((\(s, is). length init) x))" proof (rule ballI, goal_cases) case (1 x) from 1 config_rand_length have f1: "length (fst x) = length init" by fastforce from 1 config_rand_set have 2: "set (fst x) = set init" by fastforce from qs 2 have "(qs!n) \ set (fst x)" by auto then have "(index (fst x) (qs ! n)) < length init" apply(rule index_less) using f1 by auto then show ?case by (simp add: split_def) qed have "E(bind_pmf ?D (\(s, is). return_pmf (real (Suc (index s (qs ! n)))))) = E(map_pmf (\(s, is). real (Suc (index s (qs ! n)))) ?D)" by(simp add: split_def map_pmf_def) also have "\ \ E(map_pmf (\(s, is). length init) ?D)" apply(rule E_mono3) apply(fact integr_index) apply(simp) using absch by auto also have "\ = length init" by(simp add: split_def) finally show ?case by(simp add: map_pmf_def bind_assoc_pmf bind_return_pmf split_def) qed lemma T_BIT_ub: "\i set init \ T_BIT n \ n * size init" proof(induction n) case 0 show ?case by(simp add: T_BIT_def) next case (Suc n) thus ?case using t_BIT_ub[where n="n"] by (simp add: T_BIT_def) qed subsubsection "Main Lemma" lemma myub: "n < length qs \ t_BIT n + Phi(n + 1) - Phi n \ (7 / 4) * t_A n - 3/4" proof - assume nqs: "n < length qs" have "t_BIT n + Phi (n+1) - Phi n \ (7 / 4) * t_A n - 3/4" proof (cases "length init > 0") case False show ?thesis proof - from False have qsn: "(qs!n) \ set init" by auto from False have l0: "length init = 0" by auto then have "length (swaps (paid_A ! n) (s_A n)) = 0" using length_s_A by auto with l0 have 4: "t_A n = 1 + length (paid_A ! n)" unfolding t_A_def c_A_def p_A_def by(simp) have 1: "t_BIT n \ 1" using t_BIT_ub2[OF qsn] l0 by auto { fix m have "phi m = (\(b,(a,i)). phi m (b,(a,i)))" by auto also have "\ = (\(b,(a,i)). 0)" by(simp only: phi_empty2[OF l0]) finally have "phi m= (\(b,(a,i)). 0)". } note phinull=this have 2: "PhiPlus n = 0" unfolding PhiPlus_def apply(simp) apply(simp only: phinull) by (auto simp: split_def) have 3:"Phi n = 0" unfolding Phi_def apply(simp only: phinull) by (auto simp: split_def) have "t_A n \ 1 \ 1 \ 7 / 4 * (t_A n) - 3 / 4" by(simp) with 4 have 5: "1 \ 7 / 4 * (t_A n) - 3 / 4" by auto from 1 2 3 have "t_BIT n + PhiPlus n - Phi n \ 1" by auto also from 5 have "\ \ 7 / 4 * (t_A n) - 3 / 4" by auto finally show ?thesis using PhiPlus_is_Phi_Suc nqs by auto qed next case True let ?l = "length init" from True obtain l' where lSuc: "?l = Suc l'" by (metis Suc_pred) have 31: "n < length paid_A" using nqs by auto define q where "q = qs!n" define D where [simp]: "D = (config'' (BIT_init, BIT_step) qs init n)" define cost where [simp]: "cost = (\(s, is).(t s q (if (fst is) ! (index (snd is) q) then 0 else length s, [])))" define \\<^sub>2 where [simp]: "\\<^sub>2 = (\(s, is). ((phi (Suc n)) (step s q (if (fst is) ! (index (snd is) q) then 0 else length s, []),(flip (index (snd is) q) (fst is), snd is))))" define \\<^sub>0 where [simp]: "\\<^sub>0 = phi n" have inEreinziehn: "t_BIT n + Phi (n+1) - Phi n = E (map_pmf (\x. (cost x) + (\\<^sub>2 x) - (\\<^sub>0 x)) D)" proof - have "bind_pmf D (\(s, is). bind_pmf (BIT_step (s, is) (q)) (\(a,nis). return_pmf (real(t s (q) a)))) = bind_pmf D (\(s, is). return_pmf (t s q (if (fst is) ! (index (snd is) q) then 0 else length s, [])))" unfolding BIT_step_def apply (auto simp: bind_return_pmf split_def) by (metis prod.collapse) also have "\ = map_pmf cost D" by (auto simp: map_pmf_def split_def) finally have rightform1: "bind_pmf D (\(s, is). bind_pmf (BIT_step (s, is) (q)) (\(a,nis). return_pmf (real(t s (q) a)))) = map_pmf cost D" . have rightform2: "map_pmf (phi (Suc n)) (bind_pmf D (\(s, is). bind_pmf (BIT_step (s, is) (q)) (\(a, nis). return_pmf (step s (q) a, nis)))) = map_pmf \\<^sub>2 D" apply(simp add: bind_return_pmf bind_assoc_pmf map_pmf_def split_def BIT_step_def) by (metis prod.collapse) have "t_BIT n + Phi (n+1) - Phi n = t_BIT n + PhiPlus n - Phi n" using PhiPlus_is_Phi_Suc nqs by auto also have "\ = T_on_rand_n BIT init qs n + E (map_pmf (phi (Suc n)) (bind_pmf D (\(s, is). bind_pmf (BIT_step (s, is) (q)) (\(a, nis). return_pmf (step s (q) a, nis))))) - E (map_pmf (phi n) D) " unfolding PhiPlus_def Phi_def t_BIT_def q_def by auto also have "\ = E (bind_pmf D (\(s, is). bind_pmf (BIT_step (s, is) (q)) (\(a,nis). return_pmf (t s (q) a)))) + E (map_pmf (phi (Suc n)) (bind_pmf D (\(s, is). bind_pmf (BIT_step (s, is) (q)) (\(a, nis). return_pmf (step s (q) a, nis))))) - E (map_pmf \\<^sub>0 D)" by (auto simp: q_def split_def) also have "\ = E (map_pmf cost D) + E (map_pmf \\<^sub>2 D) - E (map_pmf \\<^sub>0 D)" using rightform1 rightform2 split_def by auto also have "\ = E (map_pmf (\x. (cost x) + (\\<^sub>2 x)) D) - E (map_pmf (\x. (\\<^sub>0 x)) D)" unfolding D_def using E_linear_plus2[OF finite_config_BIT[OF dist_init]] by auto also have "\ = E (map_pmf (\x. (cost x) + (\\<^sub>2 x) - (\\<^sub>0 x)) D)" unfolding D_def by(simp only: E_linear_diff2[OF finite_config_BIT[OF dist_init]] split_def) finally show "t_BIT n + Phi (n+1) - Phi n = E (map_pmf (\x. (cost x) + (\\<^sub>2 x) - (\\<^sub>0 x)) D)" by auto qed define xs where [simp]: "xs = s_A n" define xs' where [simp]: "xs' = swaps (paid_A!n) xs" define xs'' where [simp]: "xs'' = mtf2 (free_A!n) (q) xs'" define k where [simp]: "k = index xs' q" (* position of the requested element in A's list *) define k' where [simp]: "k' = max 0 (k-free_A!n)" (* position where A moves the requested element to *) have [simp]: "length xs = length init" by auto have dp_xs_init[simp]: "dist_perm xs init" by auto text "The Transformation" have ub_cost: "\x\set_pmf D. (real (cost x)) + (\\<^sub>2 x) - (\\<^sub>0 x) \ k + 1 + (if (q) \ set init then (if (fst (snd x))!(index init q) then k-k' else (\ji<(length (paid_A!n)). (if (fst (snd x))!(gebub n i) then 2 else 1))" proof (rule, goal_cases) case (1 x) note xinD=1 then have [simp]: "snd (snd x) = init" using D_def config_n_init3 by fast define b where "b = fst (snd x)" define ys where "ys = fst x" define aBIT where [simp]: "aBIT = (if b ! (index (snd (snd x)) q) then 0 else length ys, ([]::nat list))" define ys' where "ys' = step ys (q) aBIT" define b' where "b' = flip (index init q) b" define \\<^sub>1 where "\\<^sub>1 = (\z:: 'a list\ (bool list \ 'a list) . (\(x,y)\(Inv ys xs'). (if fst (snd z)!(index init y) then 2::real else 1)))" have xs''_step: "xs'' = step xs (q) (free_A!n,paid_A!n)" unfolding xs'_def xs''_def xs_def step_def free_A_def paid_A_def by(auto simp: split_def) have gis2: "(\\<^sub>2 (ys,(b,init))) = (\(x,y)\(Inv ys' xs''). (if b'!(index init y) then 2 else 1))" apply(simp only: split_def) apply(simp only: xs''_step) apply(simp only: \\<^sub>2_def phi.simps) unfolding b'_def b_def ys'_def aBIT_def q_def unfolding s_A.simps apply(simp only: split_def) by auto then have gis: "\\<^sub>2 x = (\(x,y)\(Inv ys' xs''). (if b'!(index init y) then 2 else 1))" unfolding ys_def b_def by (auto simp: split_def) have his2: "(\\<^sub>0 (ys,(b,init))) = (\(x,y)\(Inv ys xs). (if b!(index init y) then 2 else 1))" apply(simp only: split_def) apply(simp only: \\<^sub>0_def phi.simps) by(simp add: split_def) then have his: "(\\<^sub>0 x) = (\(x,y)\(Inv ys xs). (if b!(index init y) then 2 else 1))" by(auto simp: ys_def b_def split_def phi') have dis: "\\<^sub>1 x = (\(x,y)\(Inv ys xs'). (if b!(index init y) then 2 else 1))" unfolding \\<^sub>1_def b_def by auto have "ys' = mtf2 (fst aBIT) (q) ys" by (simp add: step_def ys'_def) from config_rand_distinct[of BIT] config_rand_set[of BIT] xinD have dp_ys_init[simp]: "dist_perm ys init" unfolding D_def ys_def by force have dp_ys'_init[simp]: "dist_perm ys' init" unfolding ys'_def step_def by (auto) then have lenys'[simp]: "length ys' = length init" by (metis distinct_card) have dp_xs'_init[simp]: "dist_perm xs' init" by auto have gra: "dist_perm ys xs'" by auto have leninitb[simp]: "length b = length init" using b_def config_n_fst_init_length2 xinD[unfolded] by auto have leninitys[simp]: "length ys = length init" using dp_ys_init by (metis distinct_card) {fix m have "dist_perm ys (s'_A n m)" using dp_ys_init by auto } note dist=this text "Upper bound of the inversions created by paid exchanges of A" (* ============================================ first we adress the paid exchanges paid cost of A: p_A *) let ?paidUB="(\i<(length (paid_A!n)). (if b!(gebub n i) then 2::real else 1))" have paid_ub: "\\<^sub>1 x \ \\<^sub>0 x + ?paidUB" proof - have a: "length (paid_A ! n) \ length (paid_A ! n)" by auto have b: "xs' = (s'_A n (length (paid_A ! n)))" using s'A_m by auto { fix m have "m\length (paid_A!n) \ (\(x,y)\(Inv ys (s'_A n m)). (if b!(index init y) then 2::real else 1)) \ (\(x,y)\(Inv ys xs). (if b!(index init y) then 2 else 1)) + (\i length (paid_A ! n)" and m_bd: "m < length (paid_A ! n)" by auto note yeah = Suc(1)[OF m_bd2] let ?revm="(length (paid_A ! n) - Suc m)" note ah=Inv_swap[of "ys" "(s'_A n m)" "(paid_A ! n ! ?revm)", OF dist] have "(\(xa, y)\Inv ys (s'_A n (Suc m)). if b ! (index init y) then 2::real else 1) = (\(xa, y)\Inv ys (swap (paid_A ! n ! ?revm) (s'_A n m)). if b ! (index init y) then 2 else 1)" using s'_A.simps(2) by auto also have "\ = (\(xa, y)\(if Suc (paid_A ! n ! ?revm) < length ys then if s'_A n m ! (paid_A ! n ! ?revm) < s'_A n m ! Suc (paid_A ! n ! ?revm) in ys then Inv ys (s'_A n m) \ {(s'_A n m ! (paid_A ! n ! ?revm), s'_A n m ! Suc (paid_A ! n ! ?revm))} else Inv ys (s'_A n m) - {(s'_A n m ! Suc (paid_A ! n ! ?revm), s'_A n m ! (paid_A ! n ! ?revm))} else Inv ys (s'_A n m)). if b ! (index init y) then 2::real else 1)" by (simp only: ah) also have "\ \ (\(xa, y)\Inv ys (s'_A n m). if b ! (index init y) then 2::real else 1) + (if (b) ! (index init (s'_A n m ! Suc (paid_A ! n ! ?revm))) then 2::real else 1)" (is "?A \ ?B") proof(cases "Suc (paid_A ! n ! ?revm) < length ys") case False (* FIXME! can't occur! because it has already been filtered out! see: then have "False" using paidAnm_inbound apply(auto) using m_bd nqs by blast *) then have "?A = (\(xa, y)\(Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1)" by auto also have "\ \ (\(xa, y)\(Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1) + (if b ! (index init (s'_A n m ! Suc (paid_A ! n ! ?revm))) then 2::real else 1)" by auto finally show "?A \ ?B" . next case True then have "?A = (\(xa, y)\(if s'_A n m ! (paid_A ! n ! ?revm) < s'_A n m ! Suc (paid_A ! n ! ?revm) in ys then Inv ys (s'_A n m) \ {(s'_A n m ! (paid_A ! n ! ?revm), s'_A n m ! Suc (paid_A ! n ! ?revm))} else Inv ys (s'_A n m) - {(s'_A n m ! Suc (paid_A ! n ! ?revm), s'_A n m ! (paid_A ! n ! ?revm))} ). if b ! (index init y) then 2 else 1)" by auto also have "\ \ ?B" (is "?A' \ ?B") proof (cases "s'_A n m ! (paid_A ! n ! ?revm) < s'_A n m ! Suc (paid_A ! n ! ?revm) in ys") case True let ?neurein="(s'_A n m ! (paid_A ! n ! ?revm), s'_A n m ! Suc (paid_A ! n ! ?revm))" from True have "?A' = (\(xa, y)\(Inv ys (s'_A n m) \ {?neurein} ). if b ! (index init y) then 2 else 1)" by auto also have "\ = (\(xa, y)\insert ?neurein (Inv ys (s'_A n m) ). if b ! (index init y) then 2 else 1)" by auto also have "\ \ (if b ! (index init (snd ?neurein)) then 2 else 1) + (\(xa, y)\(Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1)" proof (cases "?neurein \ Inv ys (s'_A n m)") case True then have "insert ?neurein (Inv ys (s'_A n m)) = (Inv ys (s'_A n m))" by auto then have "(\(xa, y)\insert ?neurein (Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1) = (\(xa, y)\(Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1)" by auto also have "\ \ (if b ! (index init (snd ?neurein)) then 2::real else 1) + (\(xa, y)\(Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1)" by auto finally show ?thesis . next case False have "(\(xa, y)\insert ?neurein (Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1) = (\y\insert ?neurein (Inv ys (s'_A n m)). (\i. if b ! (index init (snd i)) then 2 else 1) y)" by(auto simp: split_def) also have "\ = (\i. if b ! (index init (snd i)) then 2 else 1) ?neurein + (\y\(Inv ys (s'_A n m)) - {?neurein}. (\i. if b ! (index init (snd i)) then 2 else 1) y)" apply(rule sum.insert_remove) by(auto) also have "\ = (if b ! (index init (snd ?neurein)) then 2 else 1) + (\y\(Inv ys (s'_A n m)). (\i. if b ! (index init (snd i)) then 2::real else 1) y)" using False by auto also have "\ \ (if b ! (index init (snd ?neurein)) then 2 else 1) + (\(xa, y)\(Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1)" by(simp only: split_def) finally show ?thesis . qed also have "\ = (\(xa, y)\Inv ys (s'_A n m). if b ! (index init y) then 2 else 1) + (if b ! (index init (s'_A n m ! Suc (paid_A ! n ! ?revm))) then 2 else 1)" by auto finally show ?thesis . next case False then have "?A' = (\(xa, y)\(Inv ys (s'_A n m) - {(s'_A n m ! Suc (paid_A ! n ! ?revm), s'_A n m ! (paid_A ! n ! ?revm))} ). if b ! (index init y) then 2 else 1)" by auto also have "\ \ (\(xa, y)\(Inv ys (s'_A n m)). if b ! (index init y) then 2 else 1)" (is "(\(xa, y)\?X-{?x}. ?g y) \ (\(xa, y)\?X. ?g y) ") proof (cases "?x \ ?X") case True have "(\(xa, y)\?X-{?x}. ?g y) \ (%(xa,y). ?g y) ?x + (\(xa, y)\?X-{?x}. ?g y)" by simp also have "\ = (\(xa, y)\?X. ?g y)" apply(rule sum.remove[symmetric]) apply simp apply(fact) done finally show ?thesis . qed simp also have "\ \ ?B" by auto finally show ?thesis . qed finally show "?A \ ?B" . qed also have "\ \ (\(xa, y)\Inv ys (s_A n). if b ! (index init y) then 2::real else 1) + (\i = (\(xa, y)\Inv ys (s_A n). if b ! (index init y) then 2::real else 1) + (\i = (\(xa, y)\Inv ys (s_A n). if b ! (index init y) then 2::real else 1) + (\i<(Suc m). if b ! gebub n i then 2 else 1)" by auto finally show ?case by simp qed (simp add: split_def) } note x = this[OF a] show ?thesis unfolding \\<^sub>1_def his apply(simp only: b) using x b_def by auto qed text "Upper bound for the costs of BIT" define inI where [simp]: "inI = InvOf (q) ys xs'" define I where [simp]: "I = card(InvOf (q) ys xs')" (* ys is BITs list, xs' is A's list after paid exchanges *) have ub_cost_BIT: "(cost x) \ k + 1 + I" proof (cases "(q) \ set init") case False (* cannot occur! ! ! OBSOLETE *) from False have 4: "I = 0" by(auto simp: before_in_def) have "(cost x) = 1 + index ys (q)" by (auto simp: ys_def t_def split_def) also have "\ = 1 + length init" using False by auto also have "\ = 1 + k" using False by auto finally show ?thesis using 4 by auto next case True then have gra2: "(q) \ set ys" using dp_ys_init by auto have "(cost x) = 1 + index ys (q)" by(auto simp: ys_def t_def split_def) also have "\ \ k + 1 + I" using numberofIsbeschr[OF gra gra2] by auto finally show"(cost x) \ k + 1 + I" . qed text "Upper bound for inversions generated by free exchanges" (* ================================================ *) (* ================================================ *) (* second part: FREE EXCHANGES *) define ub_free where "ub_free = (if (q \ set init) then (if b!(index init q) then k-k' else (\j(x,y)\(Inv ys' xs''). (if b' !(index init y) then 2 else 1 ) ) - (\(x,y)\(Inv ys xs'). (if b!(index init y) then 2 else 1) ) \ ?ub2" proof (cases "(q) \ set init") case False from False have 1: "ys' = ys" unfolding ys'_def step_def mtf2_def by(simp) from False have 2: "xs' = xs''" unfolding xs''_def mtf2_def by(simp) from False have "(index init q) \ length b" using setinit by auto then have 3: "b' = b" unfolding b'_def using flip_out_of_bounds by auto from False have 4: "I = 0" unfolding I_def before_in_def by(auto) note ubnn=False have nn: "k-k'\0" unfolding k_def k'_def by auto from 1 2 3 4 have "(\(x,y)\(Inv ys' xs''). (if b'!(index init y) then 2::real else 1)) - (\(x,y)\(Inv ys xs'). (if b!(index init y) then 2 else 1)) = -I" by auto with ubnn show ?thesis unfolding ub_free_def by auto next case True note queryinlist=this then have gra2: "q \ set ys" using dp_ys_init by auto have k_inbounds: "k < length init" using index_less_size_conv queryinlist by (simp) { fix y e fix X::"bool list" assume rd: "e < length X" have "y < length X \ (if flip e X ! y then 2::real else 1) - (if X ! y then 2 else 1) = (if e=y then (if X ! y then -1 else 1) else 0)" proof cases assume "y < length X" and ey: "e=y" then have "(if flip e X ! y then 2::real else 1) - (if X ! y then 2 else 1) = (if X ! y then 1::real else 2) - (if X ! y then 2 else 1)" using flip_itself by auto also have "\ = (if X ! y then -1::real else 1)" by auto finally show "(if flip e X ! y then 2::real else 1) - (if X ! y then 2 else 1) = (if e=y then (if X ! y then -1 else 1) else 0)" using ey by auto next assume len: "y < length X" and eny: "e\y" then have "(if flip e X ! y then 2::real else 1) - (if X ! y then 2 else 1) = (if X ! y then 2::real else 1) - (if X ! y then 2 else 1)" using flip_other[OF len rd eny] by auto also have "\ = 0" by auto finally show "(if flip e X ! y then 2::real else 1) - (if X ! y then 2 else 1) = (if e=y then (if X ! y then -1 else 1) else 0)" using eny by auto qed } note flipstyle=this from queryinlist setinit have qsfst: "(index init q) < length b" by simp have fA: "finite (Inv ys' xs'')" by auto have fB: "finite (Inv ys xs')" by auto define \ where [simp]: "\ = (\(x,y)\(Inv ys' xs''). (if b'!(index init y) then 2::real else 1)) - (\(x,y)\(Inv ys xs'). (if b!(index init y) then 2 else 1))" define C where [simp]: "C = (\(x,y)\(Inv ys' xs'') \ (Inv ys xs'). (if b'!(index init y) then 2::real else 1) - (if b!(index init y) then 2 else 1))" define A where [simp]: "A = (\(x,y)\(Inv ys' xs'')-(Inv ys xs'). (if b'!(index init y) then 2::real else 1))" define B where [simp]: "B = (\(x,y)\(Inv ys xs')-(Inv ys' xs''). (if b!(index init y) then 2::real else 1))" have teilen: "\ = C + A - B" (* C A B *) unfolding \_def A_def B_def C_def using sum_my[OF fA fB] by (auto simp: split_def) then have "\ = A - B + C" by auto then have teilen2: "\\<^sub>2 x - \\<^sub>1 x = A - B + C" unfolding \_def using dis gis by auto have setys': "(index init) ` (set ys') = {0.. = {0.. = {0.. -I" proof (cases "b!(index init q)") (* case distinction on whether the bit of the requested element is set *) case True then have samesame: "ys' = ys" unfolding ys'_def step_def by auto then have puh: "(Inv ys' xs') = (Inv ys xs')" by auto { fix \ \ assume "(\,\)\(Inv ys' xs'') \ (Inv ys' xs')" then have "(\,\)\(Inv ys' xs'')" by auto then have "(\< \ in ys')" unfolding Inv_def by auto then have 1: "\ \ set ys'" by (simp only: before_in_setD2) then have "index init \ < length ys'" using setys' by auto then have "index init \ < length init" using lenys' by auto then have puzzel: "index init \ < length b" using leninitb by auto have betainit: "\ \ set init" using 1 by auto have aha: "(q=\) = (index init q = index init \)" using betainit by simp have "(if b'!(index init \) then 2::real else 1) - (if b!(index init \) then 2 else 1) = (if (index init q) = (index init \) then if b !(index init \) then - 1 else 1 else 0)" unfolding b'_def apply(rule flipstyle) by(fact)+ also have "\ = (if (index init q) = (index init \) then if b ! (index init q) then - 1 else 1 else 0)" by auto also have "\ = (if q = \ then - 1 else 0)" using aha True by auto finally have "(if b'!(index init \) then 2::real else 1) - (if b!(index init \) then 2 else 1) = (if (q) = \ then -1::real else 0)" by auto } then have grreeeaa: "\x\(Inv ys' xs'') \ (Inv ys' xs'). (\x. (if b'! (index init (snd x)) then 2::real else 1) - (if b! (index init (snd x)) then 2 else 1)) x = (\x. (if (q) = snd x then -1::real else 0)) x" by force let ?fin="(Inv ys' xs'') \ (Inv ys' xs')" have ttt: "{(x,y). (x,y)\(Inv ys' xs'') \ (Inv ys' xs') \ y = (q)} \ {(x,y). (x,y)\(Inv ys' xs'') \ (Inv ys' xs') \ y \ (q)} = (Inv ys' xs'') \ (Inv ys' xs')" (is "?split1 \ ?split2 = ?easy") by auto have interem: "?split1 \ ?split2 = {}" by auto have split1subs: "?split1 \ ?fin" by auto have split2subs: "?split2 \ ?fin" by auto have fs1: "finite ?split1" apply(rule finite_subset[where B="?fin"]) apply(rule split1subs) by(auto) have fs2: "finite ?split2" apply(rule finite_subset[where B="?fin"]) apply(rule split2subs) by(auto) have "k - k' \ (free_A!n)" by auto have g: "InvOf (q) ys' xs'' \ InvOf (q) ys' xs'" using True apply(auto) apply(rule mtf2_mono[of "swaps (paid_A ! n) (s_A n)"]) by (auto simp: queryinlist) have h: "?split1 = (InvOf (q) ys' xs'') \ (InvOf (q) ys' xs')" unfolding Inv_def by auto also from g have "\ = InvOf (q) ys' xs'" by force also from samesame have "\ = InvOf (q) ys xs'" by simp finally have "?split1 = inI" unfolding inI_def . then have cardsp1isI: "card ?split1 = I" by auto { fix a b assume "(a,b)\?split1" then have "b = (q)" by auto then have "(if (q) = b then (-1::real) else 0) = (-1::real)" by auto } then have split1easy: "\x\?split1. (\x. (if (q) = snd x then (-1::real) else 0)) x = (\x. (-1::real)) x" by force { fix a b assume "(a,b)\?split2" then have "~ b = (q)" by auto then have "(if (q) = b then (-1::real) else 0) = 0" by auto } then have split2easy: "\x\?split2. (\x. (if (q) = snd x then (-1::real) else 0)) x = (\x. 0::real) x" by force have E0: "C = (\(x,y)\(Inv ys' xs'') \ (Inv ys xs'). (if b'!(index init y) then 2::real else 1) - (if b!(index init y) then 2 else 1))" by auto also from puh have E1: "... = (\(x,y)\(Inv ys' xs'') \ (Inv ys' xs'). (if b'!(index init y) then 2::real else 1) - (if b!(index init y) then 2 else 1))" by auto also have E2: "\ = (\(x,y)\?easy. (if (q) = y then (-1::real) else 0))" using sum_my2[OF grreeeaa] by (auto simp: split_def) also have E3: "\ = (\(x,y)\?split1 \ ?split2. (if (q) = y then (-1::real) else 0))" by(simp only: ttt) also have "\ = (\(x,y)\?split1. (if (q) = y then (-1::real) else 0)) + (\(x,y)\?split2. (if (q) = y then (-1::real) else 0)) - (\(x,y)\?split1 \ ?split2. (if (q) = y then (-1::real) else 0))" by(rule sum_Un[OF fs1 fs2]) also have "\ = (\(x,y)\?split1. (if (q) = y then (-1::real) else 0)) + (\(x,y)\?split2. (if (q) = y then (-1::real) else 0))" apply(simp only: interem) by auto also have E4: "\ = (\(x,y)\?split1. (-1::real) ) + (\(x,y)\?split2. 0)" using sum_my2[OF split1easy]sum_my2[OF split2easy] by(simp only: split_def) also have "\ = (\(x,y)\?split1. (-1::real) )" by auto also have E5: "\ = - card ?split1 " by auto also have E6: "\ = - I " using cardsp1isI by auto finally have abschC: "C = -I". have abschB: "B \ (0::real)" unfolding B_def apply(rule sum_nonneg) by auto from abschB abschC show "C - B \ -I" by simp next case False from leninitys False have ya: "ys' = mtf2 (length ys) q ys" unfolding step_def ys'_def by(auto) have "index ys' q = 0" unfolding ya apply(rule mtf2_moves_to_front) using gra2 by simp_all then have nixbefore: "before q ys' = {}" unfolding before_in_def by auto { fix \ \ assume "(\,\)\(Inv ys' xs'') \ (Inv ys xs')" then have "(\,\)\(Inv ys' xs'')" by auto then have "(\< \ in ys')" unfolding Inv_def by auto then have 1: "\ \ set ys'" by (simp only: before_in_setD2) then have "(index init \) < length ys'" using setys' by auto then have "(index init \) < length init" using lenys' by auto then have puzzel: "(index init \) < length b" using leninitb by auto have betainit: "\ \ set init" using 1 by auto have aha: "(q=\) = (index init q = index init \)" using betainit by simp have "(if b'!(index init \) then 2::real else 1) - (if b!(index init \) then 2 else 1) = (if (index init q) = (index init \) then if b ! (index init \) then - 1 else 1 else 0)" unfolding b'_def apply(rule flipstyle) by(fact)+ also have "\ = (if (index init q) = (index init \) then if b ! (index init q) then - 1 else 1 else 0)" by auto also have "\ = (if (q) = \ then 1 else 0)" using False aha by auto finally have "(if b'!(index init \) then 2::real else 1) - (if b!(index init \) then 2 else 1) = (if (q) = \ then 1::real else 0)" by auto } then have grreeeaa2: "\x\(Inv ys' xs'') \ (Inv ys xs'). (\x. (if b'! (index init (snd x)) then 2::real else 1) - (if b! (index init (snd x)) then 2 else 1)) x = (\x. (if (q) = snd x then 1::real else 0)) x" by force let ?fin="(Inv ys' xs'') \ (Inv ys xs')" have ttt: "{(x,y). (x,y)\(Inv ys' xs'') \ (Inv ys xs') \ y = (q)} \ {(x,y). (x,y)\(Inv ys' xs'') \ (Inv ys xs') \ y \ (q)} = (Inv ys' xs'') \ (Inv ys xs')" (is "?split1 \ ?split2 = ?easy") by auto have interem: "?split1 \ ?split2 = {}" by auto have split1subs: "?split1 \ ?fin" by auto have split2subs: "?split2 \ ?fin" by auto have fs1: "finite ?split1" apply(rule finite_subset[where B="?fin"]) apply(rule split1subs) by(auto) have fs2: "finite ?split2" apply(rule finite_subset[where B="?fin"]) apply(rule split2subs) by(auto) have split1easy : "\x\?split1. (\x. (if (q) = snd x then (1::real) else 0)) x = (\x. (1::real)) x" by force have split2easy : "\x\?split2. (\x. (if (q) = snd x then (1::real) else 0)) x = (\x. (0::real)) x" by force from nixbefore have InvOfempty: "InvOf q ys' xs'' = {}" unfolding Inv_def by auto have "?split1 = InvOf q ys' xs'' \ InvOf q ys xs'" unfolding Inv_def by auto also from InvOfempty have "\ = {}" by auto finally have split1empty: "?split1 = {}" . have "C = (\(x,y)\?easy. (if (q) = y then (1::real) else 0))" unfolding C_def by(simp only: split_def sum_my2[OF grreeeaa2]) also have "\ = (\(x,y)\?split1 \ ?split2. (if (q) = y then (1::real) else 0))" by(simp only: ttt) also have "\ = (\(x,y)\?split1. (if (q) = y then (1::real) else 0)) + (\(x,y)\?split2. (if (q) = y then (1::real) else 0)) - (\(x,y)\?split1 \ ?split2. (if (q) = y then (1::real) else 0))" by(rule sum_Un[OF fs1 fs2]) also have "\ = (\(x,y)\?split1. (if (q) = y then (1::real) else 0)) + (\(x,y)\?split2. (if (q) = y then (1::real) else 0))" apply(simp only: interem) by auto also have "\ = (\(x,y)\?split1. (1::real) ) + (\(x,y)\?split2. 0)" using sum_my2[OF split1easy] sum_my2[OF split2easy] by (simp only: split_def) also have "\ = (\(x,y)\?split1. (1::real) )" by auto also have "\ = card ?split1" by auto also have "\ = (0::real)" apply(simp only: split1empty) by auto finally have abschC: "C = (0::real)" . (* approx for B *) have ttt2: "{(x,y). (x,y)\(Inv ys xs') - (Inv ys' xs'') \ y = (q)} \ {(x,y). (x,y)\(Inv ys xs') - (Inv ys' xs'') \ y \ (q)} = (Inv ys xs') - (Inv ys' xs'')" (is "?split1 \ ?split2 = ?easy2") by auto have interem: "?split1 \ ?split2 = {}" by auto have split1subs: "?split1 \ ?easy2" by auto have split2subs: "?split2 \ ?easy2" by auto have fs1: "finite ?split1" apply(rule finite_subset[where B="?easy2"]) apply(rule split1subs) by(auto) have fs2: "finite ?split2" apply(rule finite_subset[where B="?easy2"]) apply(rule split2subs) by(auto) from False have split1easy2: "\x\?split1. (\x. (if b! (index init (snd x)) then 2::real else 1)) x = (\x. (1::real)) x" by force have "?split1 = (InvOf q ys xs') - (InvOf q ys' xs'')" unfolding Inv_def by auto also have "\ = inI" unfolding InvOfempty by auto finally have splI: "?split1 = inI" . have abschaway: "(\(x,y)\?split2. (if b!(index init y) then 2::real else 1)) \ 0" apply(rule sum_nonneg) by auto have "B = (\(x,y)\?split1 \ ?split2. (if b!(index init y) then 2::real else 1) )" unfolding B_def by(simp only: ttt2) also have "\ = (\(x,y)\?split1. (if b!(index init y) then 2::real else 1)) + (\(x,y)\?split2. (if b!(index init y) then 2::real else 1)) - (\(x,y)\?split1 \ ?split2. (if b!(index init y) then 2::real else 1))" by(rule sum_Un[OF fs1 fs2]) also have "\ = (\(x,y)\?split1. (if b!(index init y) then 2::real else 1)) + (\(x,y)\?split2. (if b!(index init y) then 2::real else 1))" apply(simp only: interem) by auto also have "\ = (\(x,y)\?split1. 1) + (\(x,y)\?split2. (if b!(index init y) then 2::real else 1))" using sum_my2[OF split1easy2] by (simp only: split_def) also have "\ = card ?split1 + (\(x,y)\?split2. (if b!(index init y) then 2::real else 1))" by auto also have "\ = I + (\(x,y)\?split2. (if b!(index init y) then 2::real else 1))" using splI by auto also have "\ \ I" using abschaway by auto finally have abschB: "B \ I" . from abschB abschC show "C - B \ -I" by auto qed (* ========================================== central! calculations for A ========================================== *) have A_absch: "A \ (if b!(index init q) then k-k' else (\j(x,y)\(Inv ys' xs'')-(Inv ys xs'). (if b'!(index init y) then 2::real else 1))" by auto have "index (mtf2 (free_A ! n) (q) (swaps (paid_A ! n) (s_A n))) (q) = (index (swaps (paid_A ! n) (s_A n)) (q) - free_A ! n)" apply(rule mtf2_q_after) using queryinlist by auto then have whatisk': "k' = index xs'' q" by auto have ss: "set ys' = set ys" by auto have ss2: "set xs' = set xs''" by auto have di: "distinct init" by auto have dys: "distinct ys" by auto have "(Inv ys' xs'')-(Inv ys xs') = {(x,y). x < y in ys' \ y < x in xs'' \ (~x < y in ys \ ~ y < x in xs')}" unfolding Inv_def by auto also have "\ = {(x,y). y\q \ x < y in ys' \ y < x in xs'' \ (~x < y in ys \ ~ y < x in xs') }" using nixbefore by blast also have "\ = {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ (~x < y in ys \ ~ y < x in xs') }" unfolding before_in_def by auto also have "\ = {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ ~x < y in ys } \ {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ ~ y < x in xs' }" by force also have "\ = {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ y < x in ys } \ {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ ~ y < x in xs' }" using before_in_setD1[where xs="ys'"] before_in_setD2[where xs="ys'"] not_before_in ss by metis also have "\ = {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ y < x in ys } \ {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ x < y in xs' }" (is "?S1 \ ?S2 = ?S1 \ ?S2'") proof - have "?S2 = ?S2'" apply(safe) proof (goal_cases) case (2 a b) from 2(5) have "~ b < a in xs'" by auto with 2(6) show "False" by auto next case (1 a b) from 1(4) have "a \ set xs'" "b \ set xs'" using before_in_setD1[where xs="xs''"] before_in_setD2[where xs="xs''"] ss2 by auto with not_before_in 1(5) have "(a < b in xs' \ a = b)" by metis with 1(1) show "a < b in xs'" by auto qed then show ?thesis by auto qed also have "\ = {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ y < x in ys } \ {(x,y). x\y \ y\q \ x < y in ys' \ ~ x < y in xs'' \ x < y in xs' }" (is "?S1 \ ?S2 = ?S1 \ ?S2'") proof - have "?S2 = ?S2'" apply(safe) proof (goal_cases) case (1 a b) from 1(4) have "~ a < b in xs''" by auto with 1(6) show "False" by auto next case (2 a b) from 2(5) have "a \ set xs''" "b \ set xs''" using before_in_setD1[where xs="xs'"] before_in_setD2[where xs="xs'"] ss2 by auto with not_before_in 2(4) have "(b < a in xs'' \ a = b)" by metis with 2(1) show "b < a in xs''" by auto qed then show ?thesis by auto qed also have "\ = {(x,y). x\y \ y\q \ x < y in ys' \ y < x in xs'' \ y < x in ys } \ {}" using x_stays_before_y_if_y_not_moved_to_front[where xs="xs'" and q="q"] before_in_setD1[where xs="xs'"] before_in_setD2[where xs="xs'"] by (auto simp: queryinlist) also have "\ = {(x,y). x\y \ x=q \ y\q \ x < y in ys' \ y < x in xs'' \ y < x in ys }" apply(simp only: ya) using swapped_by_mtf2[where xs="ys" and q="q" and n="(length ys)"] dys before_in_setD1[where xs="ys"] before_in_setD2[where xs="ys"] by (auto simp: queryinlist) also have "\ \ {(x,y). x=q \ y\q \ q < y in ys' \ y < q in xs''}" by force also have "\ = {(x,y). x=q \ y\q \ q < y in ys' \ y < q in xs'' \ y \ set xs''}" using before_in_setD1 by metis also have "\ = {(x,y). x=q \ y\q \ q < y in ys' \ index xs'' y < index xs'' q \ q \ set xs'' \ y \ set xs''}" unfolding before_in_def by auto also have "\ = {(x,y). x=q \ y\q \ q < y in ys' \ index xs'' y < index xs' q - (free_A ! n) \ q \ set xs'' \ y \ set xs''}" using mtf2_q_after[where A="xs'" and q="q"] by force also have "\ \ {(x,y). x=q \ y\q \ index xs' y < index xs' q - (free_A ! n) \ y \ set xs''}" using mtf2_backwards_effect4'[where xs="xs'" and q="q" and n="(free_A ! n)", simplified ] by auto also have "\ \ {(x,y). x=q \ y\q \ index xs' y < k'}" using mtf2_q_after[where A="xs'" and q="q"] by auto finally have subsa: "(Inv ys' xs'')-(Inv ys xs') \ {(x,y). x=q \ y\q \ index xs' y < k'}" . have k'xs': "k' < length xs''" unfolding whatisk' apply(rule index_less) by (auto simp: queryinlist) then have k'xs': "k' < length xs'" by auto have "{(x,y). x=q \ index xs' y < k'} \ {(x,y). x=q \ index xs' y < length xs'}" using k'xs' by auto also have "\ = {(x,y). x=q \ y \ set xs'}" using index_less_size_conv by fast finally have "{(x,y). x=q \ index xs' y < k'} \ {(x,y). x=q \ y \ set xs'}" . then have finia2: "finite {(x,y). x=q \ index xs' y < k'}" apply(rule finite_subset) by(simp) have lulae: "{(a,b). a=q \ index xs' b < k'} = {(q,b)|b. index xs' b < k'}" by auto have k'b: "k' < length b" using whatisk' by (auto simp: queryinlist) have asdasd: "{(\,\). \=q \ \\q \ index xs' \ < k'} = {(\,\). \=q \ \\q \ index xs' \ < k' \ (index init \) < length b }" proof (auto, goal_cases) case (1 b) from 1(2) have "index xs' b < index xs' (q)" by auto also have "\ < length xs'" by (auto simp: queryinlist) finally have "b \ set xs'" using index_less_size_conv by metis then show ?case using setinit by auto qed { fix \ have "\\q \ (index init \)\(index init q)" using queryinlist by auto } note ij=this have subsa2: "{(\,\). \=q \ \\q \ index xs' \ < k'} \ {(\,\). \=q \ index xs' \ < k'}" by auto then have finia: "finite {(x,y). x=q \ y\q \ index xs' y < k'}" apply(rule finite_subset) using finia2 by auto have E0: "A = (\(x,y)\(Inv ys' xs'')-(Inv ys xs'). (if b'!(index init y) then 2::real else 1))" by auto also have E1: "\ \ (\(x,y)\{(a,b). a=q \ b\q \ index xs' b < k'}. (if b'!(index init y) then 2::real else 1))" unfolding A_def apply(rule sum_mono2[OF finia subsa]) by auto also have "\ = (\(x,y)\{(\,\). \=q \ \\q \ index xs' \ < k' \ (index init \) < length b }. (if b'!(index init y) then 2::real else 1))" using asdasd by auto also have "\ = (\(x,y)\{(\,\). \=q \ \\q \ index xs' \ < k' \ (index init \) < length b }. (if b!(index init y) then 2::real else 1))" proof (rule sum.cong, goal_cases) case (2 z) then obtain \ \ where zab: "z=(\, \)" and "\ = q" and diff: "\ \ q" and "index xs' \ < k'" and i: "index init \ < length b" by auto from diff ij have "index init \ \ index init q" by auto with flip_other qsfst i have "b' ! index init \ = b ! index init \" unfolding b'_def by auto with zab show ?case by(auto simp add: split_def) qed simp also have E1a: "\ = (\(x,y)\{(a,b). a=q \ b\q \ index xs' b < k'}. (if b!(index init y) then 2::real else 1))" using asdasd by auto also have "\ \ (\(x,y)\{(a,b). a=q \ index xs' b < k'}. (if b!(index init y) then 2::real else 1))" apply(rule sum_mono2[OF finia2 subsa2]) by auto also have E2: "\ = (\(x,y)\{(q,b)|b. index xs' b < k'}. (if b!(index init y) then 2::real else 1))" by (simp only: lulae[symmetric]) finally have aa: "A \ (\(x,y)\{(q,b)|b. index xs' b < k'}. (if b!(index init y) then 2::real else 1))" . have sameset: "{y. index xs' y < k'} = {xs'!i | i. i < k'}" proof (safe, goal_cases) case (1 z) show ?case proof from 1(1) have "index xs' z < index (swaps (paid_A ! n) (s_A n)) (q)" by auto also have "\ < length xs'" using index_less_size_conv by (auto simp: queryinlist) finally have "index xs' z < length xs'" . then have zset: "z \ set xs'" using index_less_size_conv by metis have f1: "xs' ! (index xs' z) = z" apply(rule nth_index) using zset by auto show "z = xs' ! (index xs' z) \ (index xs' z) < k'" using f1 1(1) by auto qed next case (2 k i) from 2(1) have "i < index (swaps (paid_A ! n) (s_A n)) (q)" by auto also have "\ < length xs'" using index_less_size_conv by (auto simp: queryinlist) finally have iset: "i < length xs'" . have "index xs' (xs' ! i) = i" apply(rule index_nth_id) using iset by(auto) with 2 show ?case by auto qed have aaa23: "inj_on (\i. xs'!i) {i. i < k'}" apply(rule inj_on_nth) apply(simp) apply(simp) proof (safe, goal_cases) case (1 i) then have "i < index xs' (q)" by auto also have "\ < length xs'" using index_less_size_conv by (auto simp: queryinlist) also have "\ = length init" by auto finally show " i < length init" . qed have aa3: "{xs'!i | i. i < k'} = (\i. xs'!i) ` {i. i < k'}" by auto have aa4: "{(q,b)|b. index xs' b < k'} = (\b. (q,b)) ` {b. index xs' b < k'}" by auto have unbelievable: "{i::nat. i < k'} = {..b. (q,b)) {b. index xs' b < k'}" unfolding inj_on_def by(simp) have "(\(x,y)\{(q,b)|b. index xs' b < k'}. (if b!(index init y) then 2::real else 1)) = (\y\{y. index xs' y < k'}. (if b!(index init y) then 2::real else 1))" proof - have "(\(x,y)\{(q,b)|b. index xs' b < k'}. (if b!(index init y) then 2::real else 1)) = (\(x,y)\ (\b. (q,b)) ` {b. index xs' b < k'}. (if b!(index init y) then 2::real else 1))" using aa4 by simp also have "\ = (\z\ (\b. (q,b)) ` {b. index xs' b < k'}. (if b!(index init (snd z)) then 2::real else 1))" by (simp add: split_def) also have "\ = (\z\{b. index xs' b < k'}. (if b!(index init (snd ((\b. (q,b)) z))) then 2::real else 1))" apply(simp only: sum.reindex[OF aadad]) by auto also have "\ = (\y\{y. index xs' y < k'}. (if b!(index init y) then 2::real else 1))" by auto finally show ?thesis . qed also have "\ = (\y\{xs'!i | i. i < k'}. (if b!(index init y) then 2::real else 1))" using sameset by auto also have "\ = (\y\(\i. xs'!i) ` {i. i < k'}. (if b!(index init y) then 2::real else 1))" using aa3 by simp also have "\ = (\y\{i::nat. i < k'}. (if b!(index init (xs'!y)) then 2::real else 1))" using sum.reindex[OF aaa23] by simp also have E3: "\ = (\j::nat(x,y)\{(q,b)|b. index xs' b < k'}. (if b!(index init y) then 2::real else 1)) = (\j (\j (if b!(index init q) then k-k' else (\j y < x in xs'' \ ~ y < x in xs'}" unfolding Inv_def using samesame by auto also have "\ \ {(xs'!i,q)|i. i\{k'.. b < a in xs'" then have anb: "a \ b" using no_before_inI by(force) have a: "a \ set init" and b: "b \ set init" using before_in_setD1[OF 1] before_in_setD2[OF 1] by auto with anb 3 have 3: "a < b in xs'" by (simp add: not_before_in) note all= anb 1 2 3 a b have bq: "b=q" apply(rule swapped_by_mtf2[where xs="xs'" and x=a]) using queryinlist apply(simp_all add: all) using all(4) apply(simp) using all(3) apply(simp) done note mine=mtf2_backwards_effect3[THEN conjunct1] from bq have "q < a in xs''" using 2 by auto then have "(k' < index xs'' a \ a \ set xs'')" unfolding before_in_def using whatisk' by auto then have low : "k' \ index xs' a" unfolding whatisk' unfolding xs''_def apply(subst mtf2_q_after) apply(simp) using queryinlist apply(simp) apply(rule mine) apply (simp add: queryinlist) using bq b apply(simp) apply(simp) apply(simp del: xs'_def) apply (metis "3" a before_in_def bq dp_xs'_init k'_def k_def max_0L mtf2_forward_beforeq nth_index whatisk' xs''_def) using a by(simp)(* unfolding xs'_def xs_def sledgehammer TODO: make this step readable by (metis "3" mtf2_q_after a before_in_def bq dp_xs'_init index_less_size_conv mtf2_forward_beforeq nth_index whatisk' xs''_def xs'_def xs_def) *) from bq have "a < q in xs'" using 3 by auto then have up: "(index xs' a < k )" unfolding before_in_def by auto from a have "a \ set xs'" by simp then have aa: "a = xs'!index xs' a" using nth_index by simp have inset: "index xs' a \ {k'.. index xs' a \ {k'.. {(xs'!i,q)|i. i\{k'.. ?UB") . have card_of_UB: "card {(xs'!i,q)|i. i\{k'.. = card ((%i. xs' ! i) ` {k'.. = card {k'.. = k-k'" by auto finally show ?thesis . qed have flipit: "flip (index init q) b ! (index init q) = (~ (b) ! (index init q))" apply(rule flip_itself) using queryinlist setinit by auto have q: "{x\?UB. snd x=q} = ?UB" by auto have E0: "A = (\(x,y)\(Inv ys' xs'')-(Inv ys xs'). (if b'!(index init y) then 2::real else 1))" by auto also have E1: "\ \ (\(z,y)\?UB. if flip (index init q) (b) ! (index init y) then 2::real else 1)" unfolding b'_def apply(rule sum_mono2[OF _ a]) by(simp_all add: split_def) also have "\ = (\(z,y)\{x\?UB. snd x=q}. if flip (index init q) (b) ! (index init y) then 2::real else 1)" by(simp only: q) also have "\ = (\z\{x\?UB. snd x=q}. if flip (index init q) (b) ! (index init (snd z)) then 2::real else 1)" by(simp add: split_def) also have "\ = (\z\{x\?UB. snd x=q}. if flip (index init q) (b) ! (index init q) then 2::real else 1)" by simp also have E2: "\ = (\z\?UB. if flip (index init q) (b) ! (index init q) then 2::real else 1)" by(simp only: q) also have E3: "\ = (\y\?UB. 1)" using flipit True by simp also have E4: "\ = k-k'" by(simp only: real_of_card[symmetric] card_of_UB) finally have result: "A \ k-k'" . with True show ?thesis by auto qed show "(\(x,y)\(Inv ys' xs''). (if b'!(index init y) then 2::real else 1)) - (\(x,y)\(Inv ys xs'). (if b!(index init y) then 2::real else 1)) \ ?ub2" unfolding ub_free_def teilen[unfolded \_def A_def B_def C_def] using BC_absch A_absch using True by auto qed from paid_ub have kl: "\\<^sub>1 x \ \\<^sub>0 x + ?paidUB" by auto from free_ub have kl2: "\\<^sub>2 x - ?ub2 \ \\<^sub>1 x" using gis dis by auto have iub_free: "I + ?ub2 = ub_free" by auto from kl kl2 have "\\<^sub>2 x - \\<^sub>0 x \ ?ub2 + ?paidUB" by auto then have "(cost x) + (\\<^sub>2 x) - (\\<^sub>0 x) \ k + 1 + I + ?ub2 + ?paidUB" using ub_cost_BIT by auto then show ?case unfolding ub_free_def b_def by auto qed text "Approximation of the Term for Free exchanges" have free_absch: "E(map_pmf (\x. (if (q) \ set init then (if (fst (snd x))!(index init q) then k-k' else (\j 3/4 * k" (is "?EA \ ?absche") proof (cases "(q) \ set init") case False then have "?EA = 0" by auto then show ?thesis by auto next case True note queryinlist=this have "k-k' \ k" by auto have "k' \ k" by auto text "Transformation of the first term" have qsn: "{index init q} \ {} \ {0.. l!(index init q)} = {xs. Ball {(index init q)} ((!) xs) \ (\i\{}. \ xs ! i) \ length xs = ?l}" by auto then have "card {l::bool list. length l = ?l \ l!(index init q)} = card {xs. Ball {index init q} ((!) xs) \ (\i\{}. \ xs ! i) \ length xs = length init} " by auto also have "\ = 2^(length init - card {index init q} - card {})" apply(subst card2[of "{(index init q)}" "{}" "?l"]) using qsn by auto finally have lulu: "card {l::bool list. length l = ?l \ l!(index init q)} = 2^(?l-1)" by auto have "(\x\{l::bool list. length l = ?l \ l!(index init q)}. real(k-k')) = (\x\{l::bool list. length l = ?l \ l!(index init q)}. k-k')" by auto also have "\ = (k-k')*2^(?l-1)" using lulu by simp finally have absch1stterm: "(\x\{l::bool list. length l = ?l \ l!(index init q)}. real(k-k')) = real((k-k')*2^(?l-1))" . text "Transformation of the second term" let ?S="{(xs'!j)|j. j set (swaps (paid_A ! n) (s_A n))" by auto then have "index (swaps (paid_A ! n) (s_A n)) q < length xs'" by auto then have k'inbound: "k' < length xs'" by auto { fix x have a: "{..jt. (if x!(index init t) then 2::real else 1)) (xs'!j)) = sum ((\t. (if x!(index init t) then 2::real else 1)) o (%j. xs'!j)) {.. = sum ((\t. (if x!(index init t) then 2::real else 1)) o (%j. xs'!j)) {j. j = sum (\t. (if x!(index init t) then 2::real else 1)) ((%j. xs'!j) ` {j. jjt. (if x!(index init t) then 2::real else 1)) (xs'!j)) = (\j\?S. (\t. (if x!(index init t) then 2 else 1)) j)" using b by simp } note reindex=this have identS: "?S = set (take k' xs')" proof - have "index (swaps (paid_A ! n) (s_A n)) (q) \ length (swaps (paid_A ! n) (s_A n))" by (rule index_le_size) then have kxs': "k' \ length xs'" by simp have "?S = (!) xs' ` {0.. = set (take k' xs')" apply(rule nth_image) by(rule kxs') finally show "?S = set (take k' xs')" . qed have distinctS: "distinct (take k' xs')" using distinct_take identS by simp have lengthS: "length (take k' xs') = k'" using length_take k'inbound by simp from distinct_card[OF distinctS] lengthS have "card (set (take k' xs')) = k'" by simp then have cardS: "card ?S = k'" using identS by simp have a: "?S \ set xs'" using set_take_subset identS by metis then have Ssubso: "(index init) ` ?S \ {0.. set init" by auto note index_inj_on_S=subset_inj_on[OF inj_on_index[of "init"] s_subst_init] have l: "xs'!k = q" unfolding k_def apply(rule nth_index) using queryinlist by(auto) have "xs'!k \ set (take k' xs')" apply(rule index_take) using l by simp then have requestnotinS: "(q) \ ?S" using l identS by simp then have indexnotin: "index init q \ (index init) ` ?S" using index_inj_on_S s_subst_init by auto have lua: "{l. length l = ?l \ ~l!(index init q)} = {xs. (\i\{}. xs ! i) \ (\i\{index init q}. \ xs ! i) \ length xs = ?l}" by auto from k'inbound have k'inbound2: "Suc k' \ length init" using Suc_le_eq by auto (* rewrite from sum over indices of the list to sum over elements (thus indices of the bit vector) *) have "(\x\{l::bool list. length l = ?l \ ~l!(index init q)}. (\jx\{l. length l = ?l \ ~l!(index init q)}. (\j\?S. (\t. (if x!(index init t) then 2 else 1)) j))" using reindex by auto (* rewrite to conform the syntax of Expactation2or1 *) also have "\ = (\x\{xs. (\i\{}. xs ! i) \ (\i\{index init q}. \ xs ! i) \ length xs = ?l}. (\j\?S. (\t. (if x!(index init t) then 2 else 1)) j))" using lua by auto also have "\ = (\x\{xs. (\i\{}. xs ! i) \ (\i\{index init q}. \ xs ! i) \ length xs = ?l}. (\j\(index init) ` ?S. (\t. (if x!t then 2 else 1)) j))" proof - { fix x have "(\j\?S. (\t. (if x!(index init t) then 2 else 1)) j) = (\j\(index init) ` ?S. (\t. (if x!t then 2 else 1)) j)" apply(simp only: sum.reindex[OF index_inj_on_S, where g="(%j. if x ! j then 2 else 1)"]) by(simp) } note a=this show ?thesis by(simp only: a) qed (* use Expactation2or1, and solve all the conditions *) also have "\ = 3 / 2 * real (card ?S) * 2 ^ (?l - card {} - card {q})" apply(subst Expactation2or1) apply(simp) apply(simp) apply(simp) apply(simp only: card_image index_inj_on_S cardS ) apply(simp add: k'inbound2 del: k'_def) - using indexnotin applysimp + using indexnotin apply simp apply(simp) using Ssubso queryinlist apply(simp) apply(simp only: card_image[OF index_inj_on_S]) by simp finally have "(\x\{l. length l = ?l \ \ l ! (index init q)}. \jx\{l. length l = ?l \ \ l ! (index init q)}. \j l!(index init q)} = 2^(?l-1)" using lulu by auto have splitie: "{l::bool list. length l = ?l} = {l::bool list. length l = ?l \ l!(index init q)} \ {l::bool list. length l = ?l \ ~l!(index init q)}" by auto have interempty: "{l::bool list. length l = ?l \ l!(index init q)} \ {l::bool list. length l = ?l \ ~l!(index init q)} = {}" by auto have fa: "finite {l::bool list. length l = ?l \ l!(index init q)}" using bitstrings_finite by auto have fb: "finite {l::bool list. length l = ?l \ ~l!(index init q)}" using bitstrings_finite by auto { fix f :: "bool list \ real" have "(\x\{l::bool list. length l = ?l}. f x) = (\x\{l::bool list. length l = ?l \ l!(index init q)} \ {l::bool list. length l = ?l \ ~l!(index init q)}. f x)" by(simp only: splitie) also have "\ = (\x\{l::bool list. length l = ?l \ l!(index init q)}. f x) + (\x\{l::bool list. length l = ?l \ ~l!(index init q)}. f x) - (\x\{l::bool list. length l = ?l \ l!(index init q)} \ {l::bool list. length l = ?l \ ~l!(index init q)}. f x)" using sum_Un[OF fa fb, of "f"] by simp also have "\ = (\x\{l::bool list. length l = ?l \ l!(index init q)}. f x) + (\x\{l::bool list. length l = ?l \ ~l!(index init q)}. f x)" by(simp add: interempty) finally have "sum f {l. length l = length init} = sum f {l. length l = length init \ l ! (index init q)} + sum f {l. length l = length init \ \ l ! (index init q)}" . } note darfstsplitten=this have E1: "E(map_pmf (\x. (if (fst (snd x))!(index init q) then real(k-k') else (\jx. (if x!(index init q) then real(k-k') else (\j snd) D))" proof - have triv: "\x. (fst \ snd) x = fst (snd x)" by simp have "E((map_pmf (\x. (if (fst (snd x))!(index init q) then real(k-k') else (\jx. ((\y. (if y!(index init q) then real(k-k') else (\j (fst \ snd)) x) D)" apply(auto simp: comp_assoc) by (simp only: triv) also have "\ = E((map_pmf (\x. (if x!(index init q) then real(k-k') else (\j (map_pmf (fst \ snd))) D)" using map_pmf_compose by metis also have "\ = E(map_pmf (\x. (if x!(index init q) then real(k-k') else (\j snd) D))" by auto finally show ?thesis . qed also have E2: "\ = E(map_pmf (\x. (if x!(index init q) then real(k-k') else (\j = (\x\(set_pmf (bv ?l)). (?insf x) * pmf (bv ?l) x)" by (subst E_finite_sum_fun) (auto simp: bv_finite mult_ac) also have "\ = (\x\{l::bool list. length l = ?l}. (?insf x) * pmf (bv ?l) x)" using bv_set by auto also have E4: "\ = (\x\{l::bool list. length l = ?l}. (?insf x) * (1/2)^?l)" by (simp add: list_pmf) also have "\ = (\x\{l::bool list. length l = ?l}. (?insf x)) * ((1/2)^?l)" by(simp only: sum_distrib_right[where r="(1/2)^?l"]) also have E5: "\ = ((1/2)^?l) *(\x\{l::bool list. length l = ?l}. (?insf x))" by(auto) also have E6: "\ = ((1/2)^?l) * ( (\x\{l::bool list. length l = ?l \ l!(index init q)}. ?insf x) + (\x\{l::bool list. length l = ?l \ ~l!(index init q)}. ?insf x) )" using darfstsplitten by auto also have E7: "\ = ((1/2)^?l) * ( (\x\{l::bool list. length l = ?l \ l!(index init q)}. ((\x. real(k-k'))) x) + (\x\{l::bool list. length l = ?l \ ~l!(index init q)}. ((\x. (\jx. (if (fst (snd x))!(index init q) then real(k-k') else (\jx\{l::bool list. length l = ?l \ l!(index init q)}. ((\x. real(k-k'))) x) + (\x\{l::bool list. length l = ?l \ ~l!(index init q)}. ((\x. (\j = ((1/2)^?l) * ( (\x\{l::bool list. length l = ?l \ l!(index init q)}. real(k-k')) + (3/2)*(real (k'))*2^(?l-1) )" by(simp only: absch2ndterm) also have E8: "\ = ((1/2)^?l) * ( real((k-k')*2^(?l-1)) + (3/2)*(real (k'))*2^(?l-1))" by(simp only: absch1stterm) (* from here it is only arithmetic ... *) also have "\ = ((1/2)^?l) * ( ( (k-k') + (k')*(3/2) ) * 2^(?l-1) )" apply(simp only: distrib_right) by simp also have "\ = ((1/2)^?l) * 2^(?l-1) * ( (k-k') + (k')*(3/2) )" by simp also have "\ = (((1::real)/2)^(Suc l')) * 2^(l') * ( real(k-k') + (k')*(3/2) )" using lSuc by auto (* REFACTOR: the only place where I use lSuc , can I avoid it? yes, if ?l=0 then k=k' = (1/2) * ( real(k-k') + (k')*(3/2) )" proof - have "((1::real)/2)^l' * 2^l' = ((1::real)/2 * 2)^l' " by(rule power_mult_distrib[symmetric]) also have "... = 1" by auto finally have "(((1::real)/2)^(Suc l'))* 2^l'=(1/2)" by auto then show ?thesis by auto qed also have E10: "\ \ (1/2) * ( (3/2)*(k-k') + (k')*(3/2) )" by auto (* and one inequality *) also have "\ = (1/2) * ( (3/2)*(k-k'+(k')) )" by auto also have "\ = (1/2) * ( (3/2)*(k) )" by auto also have E11: "\ = (3/4)*(k )" by auto finally show "E(map_pmf (\x. (if q \ set init then (if (fst (snd x))!(index init q) then real( k-k' ) else (\j 3/4 * k " using True by simp qed (* free_absch *) text "Transformation of the Term for Paid Exchanges" have paid_absch: "E(map_pmf (\x. (\i<(length (paid_A!n)). (if (fst (snd x))!(gebub n i) then 2::real else 1) )) D) = 3/2 * (length (paid_A!n))" proof - { fix i assume inbound: "(index init i) < length init" have "map_pmf (\xx. if fst (snd xx) ! (index init i) then 2::real else 1) D = bind_pmf (map_pmf (fst \ snd) D) (\b. return_pmf (if b! index init i then 2::real else 1))" unfolding map_pmf_def by(simp add: bind_assoc_pmf bind_return_pmf) also have "\ = bind_pmf (bv (length init)) (\b. return_pmf (if b! index init i then 2::real else 1))" using config_n_bv[of init "take n qs"] by simp also have "\ = map_pmf (\yy. (if yy then 2 else 1)) ( map_pmf (\y. y!(index init i)) (bv (length init)))" by (simp add: map_pmf_def bind_return_pmf bind_assoc_pmf) also have "\ = map_pmf (\yy. (if yy then 2 else 1)) (bernoulli_pmf (5 / 10))" by (auto simp add: bv_comp_bernoulli[OF inbound]) finally have "map_pmf (\xx. if fst (snd xx) ! (index init i) then 2::real else 1) D = map_pmf (\yy. if yy then 2::real else 1) (bernoulli_pmf (5 / 10)) " . } note umform = this have "E(map_pmf (\x. (\i<(length (paid_A!n)). (if (fst (snd x))!(gebub n i) then 2::real else 1))) D) = (\i<(length (paid_A!n)). E(map_pmf ((\xx. (if (fst (snd xx))!(gebub n i) then 2::real else 1))) D))" apply(subst E_linear_sum2) using finite_config_BIT[OF dist_init] by(simp_all) also have "\ = (\i<(length (paid_A!n)). E(map_pmf (\y. if y then 2::real else 1) (bernoulli_pmf (5 / 10))))" using umform gebub_def gebub_inBound[OF 31] by simp also have "\ = 3/2 * (length (paid_A!n))" by(simp add: E_bernoulli) finally show "E(map_pmf (\x. (\i<(length (paid_A!n)). (if (fst (snd x))!(gebub n i) then 2::real else 1))) D) = 3/2 * (length (paid_A!n))" . qed text "Combine the Results" (* cost of A *) have costA_absch: "k+(length (paid_A!n)) + 1 = t_A n" unfolding k_def q_def c_A_def p_A_def t_A_def by (auto) (* combine *) let ?yo= "(\x. (cost x) + (\\<^sub>2 x) - (\\<^sub>0 x))" let ?yo2=" (\x. (k + 1) + (if (q)\set init then (if (fst (snd x))!(index init q) then k-k' else (\ji<(length (paid_A!n)). (if (fst (snd x))!(gebub n i) then 2 else 1)))" have E0: "t_BIT n + Phi(n+1) - Phi n = E (map_pmf ?yo D) " using inEreinziehn by auto also have "\ \ E(map_pmf ?yo2 D)" apply(rule E_mono2) unfolding D_def apply(fact finite_config_BIT[OF dist_init]) apply(fact ub_cost[unfolded D_def]) done also have E2: "\ = E(map_pmf (\x. k + 1::real) D) + (E(map_pmf (\x. (if (q)\set init then (if (fst (snd x))!(index init q) then real(k-k') else (\jx. (\i<(length (paid_A!n)). (if (fst (snd x))!(gebub n i) then 2::real else 1))) D))" unfolding D_def apply(simp only: E_linear_plus2[OF finite_config_BIT[OF dist_init]]) by(auto simp: add.assoc) also have E3: "\ \ k + 1 + (3/4 * (real (k)) + (3/2 * real (length (paid_A!n))))" using paid_absch free_absch by auto also have "\ = k + (3/4 * (real k)) + 1 + 3/2 *(length (paid_A!n)) " by auto (* arithmetic! *) also have "\ = (1+3/4) * (real k) + 1 + 3/2 *(length (paid_A!n)) " by auto (* arithmetic! *) also have E4: "\ = 7/4*(real k) + 3/2 *(length (paid_A!n)) + 1 " by auto (* arithmetic! *) also have "\ \ 7/4*(real k) + 7/4 *(length (paid_A!n)) + 1" by auto (* arithmetic! *) also have E5:"\ = 7/4*(k+(length (paid_A!n))) + 1 " by auto also have E6:"\ = 7/4*(t_A n - (1::real)) + 1" using costA_absch by auto also have "\ = 7/4*(t_A n) - 7/4 + 1" by algebra also have E7: "\ = 7/4*(t_A n)- 3/4" by auto finally show "t_BIT n + Phi(n+1) - Phi n \ (7 / 4) * t_A n - 3/4" . qed then show "t_BIT n + Phi(n + 1) - Phi n \ (7 / 4) * t_A n - 3/4" . qed subsubsection "Lift the Result to the Whole Request List" lemma T_BIT_absch_le: assumes nqs: "n \ length qs" shows "T_BIT n \ (7 / 4) * T_A n - 3/4*n" unfolding T_BIT_def T_A_def proof - from potential2[of "Phi", OF phi0 phi_pos myub] nqs have "sum t_BIT {.. (\i = (\ii = (\ii = (\i = (7 / 4) * (\i = (7 / 4) * real_of_int (\i 7 / 4 * real_of_int (sum t_A {.. length qs" shows "T_BIT n \ (7 / 4) * T_A' n - 3/4*n" using nqs T_BIT_absch_le[of n] T_A_A'_leq[of n] by auto lemma T_A_nneg: "0 \ T_A n" by(auto simp add: sum_nonneg T_A_def t_A_def c_A_def p_A_def) lemma T_BIT_eq: "T_BIT (length qs) = T_on_rand BIT init qs" unfolding T_BIT_def T_on_rand_as_sum using t_BIT_def by auto corollary T_BIT_competitive: assumes "n \ length qs" and "init \ []" and "\i set init" shows "T_BIT n \ ((7 / 4) - 3/(4 * size init)) * T_A' n" proof cases assume 0: "real_of_int(T_A' n) \ n * (size init)" then have 1: "3/4*real_of_int(T_A' n) \ 3/4*(n * (size init))" by auto have "T_BIT n \ (7 / 4) * T_A' n - 3/4*n" using T_BIT_absch[OF assms(1)] by auto also have "\ = ((7 / 4) * real_of_int(T_A' n)) - (3/4*(n * size init)) / size init" using assms(2) by simp also have "\ \ ((7 / 4) * real_of_int(T_A' n)) - 3/4*T_A' n / size init" by(rule diff_left_mono[OF divide_right_mono[OF 1]]) simp also have "\ = ((7 / 4) - 3/4 / size init) * T_A' n" by algebra also have "\ = ((7 / 4) - 3/(4 * size init)) * T_A' n" by simp finally show ?thesis . next assume 0: "\ real_of_int(T_A' n) \ n * (size init)" have T_A'_nneg: "0 \ T_A' n" using T_A_nneg[of n] T_A_A'_leq[of n] assms(1) by auto have "2 - 1 / size init \ 1" using assms(2) by (auto simp add: field_simps neq_Nil_conv) have " T_BIT n \ n * size init" using T_BIT_ub[OF assms(3)] by linarith also have "\ < real_of_int(T_A' n)" using 0 by linarith also have "\ \ ((7 / 4) - 3/4 / size init) * T_A' n" using assms(2) T_A'_nneg by(auto simp add: mult_le_cancel_right1 field_simps neq_Nil_conv) finally show ?thesis by simp qed lemma t_A'_t: "n < length qs \ t_A' n = int (t (s_A' n) (qs!n) (acts ! n))" by (simp add: t_A'_def t_def c_A'_def p_A'_def paid_A'_def len_acts split: prod.split) lemma T_A'_eq_lem: "(\i=0.. n < length qs" thus ?case by (simp add: len_acts) qed qed lemma T_A'_eq: "T_A' (length qs) = T init qs acts" using T_A'_eq_lem by(simp add: T_A'_def atLeast0LessThan) corollary BIT_competitive3: "init \ [] \ \i set init \ T_BIT (length qs) \ ( (7/4) - 3 / (4 * length init)) * T init qs acts" using order.refl T_BIT_competitive[of "length qs"] T_A'_eq by (simp add: of_int_of_nat_eq) corollary BIT_competitive2: "init \ [] \ \i set init \ T_on_rand BIT init qs \ ( (7/4) - 3 / (4 * length init)) * T init qs acts" using BIT_competitive3 T_BIT_eq by auto corollary BIT_absch_le: "init \ [] \ T_on_rand BIT init qs \ (7 / 4) * (T init qs acts) - 3/4 * length qs" using T_BIT_absch[of "length qs", unfolded T_A'_eq T_BIT_eq] by auto end subsubsection "Generalize Competitivness of BIT" lemma setdi: "set xs = {0.. distinct xs" apply(rule card_distinct) by auto theorem compet_BIT: assumes "init \ []" "distinct init" "set qs \ set init" shows "T_on_rand BIT init qs \ ( (7/4) - 3 / (4 * length init)) * T_opt init qs" proof- from assms(3) have 1: "\i < length qs. qs!i \ set init" by auto { fix acts :: "answer list" assume len: "length acts = length qs" interpret BIT_Off acts qs init proof qed (auto simp: assms(2) len) from BIT_competitive2[OF assms(1) 1] assms(1) have "T_on_rand BIT init qs / ( (7/4) - 3 / (4 * length init)) \ real(T init qs acts)" by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) } hence "T_on_rand BIT init qs / ( (7/4) - 3 / (4 * length init)) \ T_opt init qs" apply(simp add: T_opt_def Inf_nat_def) apply(rule LeastI2_wellorder) using length_replicate[of "length qs" undefined] apply fastforce apply auto done thus ?thesis using assms by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) qed theorem compet_BIT4: assumes "init \ []" "distinct init" shows "T_on_rand BIT init qs \ 7/4 * T_opt init qs" proof- { fix acts :: "answer list" assume len: "length acts = length qs" interpret BIT_Off acts qs init proof qed (auto simp: assms(2) len) from BIT_absch_le[OF assms(1)] assms(1) have "(T_on_rand BIT init qs + 3 / 4 * length qs)/ (7/4) \ real(T init qs acts)" by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) } hence "(T_on_rand BIT init qs + 3 / 4 * length qs)/ (7/4) \ T_opt init qs" apply(simp add: T_opt_def Inf_nat_def) apply(rule LeastI2_wellorder) using length_replicate[of "length qs" undefined] apply fastforce apply auto done thus ?thesis by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) qed theorem compet_BIT_2: "compet_rand BIT (7/4) {init. init \ [] \ distinct init}" unfolding compet_rand_def proof fix init assume "init \ {init. init \ [] \ distinct init }" then have ne: "init \ []" and a: "distinct init" by auto { fix qs assume "init \ []" and a: "distinct init" then have "T_on_rand BIT init qs \ 7/4 * T_opt init qs" using compet_BIT4[of init qs] by simp } with a ne show "\b\0. \qs. static init qs \ T_on_rand BIT init qs \ (7 / 4) * (T_opt init qs) + b" by auto qed end diff --git a/thys/List_Update/Prob_Theory.thy b/thys/List_Update/Prob_Theory.thy --- a/thys/List_Update/Prob_Theory.thy +++ b/thys/List_Update/Prob_Theory.thy @@ -1,573 +1,573 @@ (* Title: Definition of Expectation and Distribution of uniformly distributed bit vectors Author: Max Haslbeck *) section "Probability Theory" theory Prob_Theory imports "HOL-Probability.Probability" begin lemma integral_map_pmf[simp]: fixes f::"real \ real" shows "(\x. f x \(map_pmf g M)) = (\x. f (g x) \M)" unfolding map_pmf_rep_eq using integral_distr[of g "(measure_pmf M)" "(count_space UNIV)" f] by auto subsection "function \E\" definition E :: "real pmf \ real" where "E M = (\x. x \ measure_pmf M)" translations "\ x. f \M" <= "CONST lebesgue_integral M (\x. f)" notation (latex output) E ("E[_]" [1] 100) lemma E_const[simp]: "E (return_pmf a) = a" unfolding E_def unfolding return_pmf.rep_eq by (simp add: integral_return) lemma E_null[simp]: "E (return_pmf 0) = 0" by auto lemma E_finite_sum: "finite (set_pmf X) \ E X = (\x\(set_pmf X). pmf X x * x)" unfolding E_def by (subst integral_measure_pmf) simp_all lemma E_of_const: "E(map_pmf (\x. y) (X::real pmf)) = y" by auto lemma E_nonneg: shows "(\x\set_pmf X. 0\ x) \ 0 \ E X" unfolding E_def using integral_nonneg by (simp add: AE_measure_pmf_iff integral_nonneg_AE) lemma E_nonneg_fun: fixes f::"'a\real" shows "(\x\set_pmf X. 0\f x) \ 0 \ E (map_pmf f X)" using E_nonneg by auto lemma E_cong: fixes f::"'a \ real" shows "finite (set_pmf X) \ (\x\ set_pmf X. (f x) = (u x)) \ E (map_pmf f X) = E (map_pmf u X)" unfolding E_def integral_map_pmf apply(rule integral_cong_AE) apply(simp add: integrable_measure_pmf_finite)+ by (simp add: AE_measure_pmf_iff) lemma E_mono3: fixes f::"'a \ real" shows " integrable (measure_pmf X) f \ integrable (measure_pmf X) u \ (\x\ set_pmf X. (f x) \ (u x)) \ E (map_pmf f X) \ E (map_pmf u X)" unfolding E_def integral_map_pmf apply(rule integral_mono_AE) by (auto simp add: AE_measure_pmf_iff) lemma E_mono2: fixes f::"'a \ real" shows "finite (set_pmf X) \ (\x\ set_pmf X. (f x) \ (u x)) \ E (map_pmf f X) \ E (map_pmf u X)" unfolding E_def integral_map_pmf apply(rule integral_mono_AE) apply(simp add: integrable_measure_pmf_finite)+ by (simp add: AE_measure_pmf_iff) lemma E_linear_diff2: "finite (set_pmf A) \ E (map_pmf f A) - E (map_pmf g A) = E (map_pmf (\x. (f x) - (g x)) A)" unfolding E_def integral_map_pmf apply(rule Bochner_Integration.integral_diff[of "measure_pmf A" f g, symmetric]) by (simp_all add: integrable_measure_pmf_finite) lemma E_linear_plus2: "finite (set_pmf A) \ E (map_pmf f A) + E (map_pmf g A) = E (map_pmf (\x. (f x) + (g x)) A)" unfolding E_def integral_map_pmf apply(rule Bochner_Integration.integral_add[of "measure_pmf A" f g, symmetric]) by (simp_all add: integrable_measure_pmf_finite) lemma E_linear_sum2: "finite (set_pmf D) \ E(map_pmf (\x. (\ii<(up::nat). E(map_pmf (f i) D))" unfolding E_def integral_map_pmf apply(rule Bochner_Integration.integral_sum) by (simp add: integrable_measure_pmf_finite) lemma E_linear_sum_allg: "finite (set_pmf D) \ E(map_pmf (\x. (\i\ A. f i x)) D) = (\i\ (A::'a set). E(map_pmf (f i) D))" unfolding E_def integral_map_pmf apply(rule Bochner_Integration.integral_sum) by (simp add: integrable_measure_pmf_finite) lemma E_finite_sum_fun: "finite (set_pmf X) \ E (map_pmf f X) = (\x\set_pmf X. pmf X x * f x)" proof - assume finite: "finite (set_pmf X)" have "E (map_pmf f X) = (\x. f x \measure_pmf X)" unfolding E_def by auto also have "\ = (\x\set_pmf X. pmf X x * f x)" by (subst integral_measure_pmf) (auto simp add: finite) finally show ?thesis . qed lemma E_bernoulli: "0\p \ p\1 \ E (map_pmf f (bernoulli_pmf p)) = p*(f True) + (1-p)*(f False)" unfolding E_def by (auto) subsection "function \bv\" fun bv:: "nat \ bool list pmf" where "bv 0 = return_pmf []" | "bv (Suc n) = do { (xs::bool list) \ bv n; (x::bool) \ (bernoulli_pmf 0.5); return_pmf (x#xs) }" lemma bv_finite: "finite (bv n)" by (induct n) auto lemma len_bv_n: "\xs \ set_pmf (bv n). length xs = n" apply(induct n) by auto lemma bv_set: "set_pmf (bv n) = {x::bool list. length x = n}" proof (induct n) case (Suc n) then have "set_pmf (bv (Suc n)) = (\x\{x. length x = n}. {True # x, False # x})" by(simp add: set_pmf_bernoulli UNIV_bool) also have "\ = {x#xs| x xs. length xs = n}" by auto also have "\ = {x. length x = Suc n} " using Suc_length_conv by fastforce finally show ?case . qed (simp) lemma len_not_in_bv: "length xs \ n \ xs \ set_pmf (bv n)" by(auto simp: len_bv_n) lemma not_n_bv_0: "length xs \ n \ pmf (bv n) xs = 0" by (simp add: len_not_in_bv pmf_eq_0_set_pmf) lemma bv_comp_bernoulli: "n < l \ map_pmf (\y. y!n) (bv l) = bernoulli_pmf (5 / 10)" proof (induct n arbitrary: l) case 0 then obtain m where "l = Suc m" by (metis Suc_pred) then show "map_pmf (\y. y!0) (bv l) = bernoulli_pmf (5 / 10)" by (auto simp: map_pmf_def bind_return_pmf bind_assoc_pmf bind_return_pmf') next case (Suc n) then have "0 < l" by auto then obtain m where lsm: "l = Suc m" by (metis Suc_pred) with Suc(2) have nltm: "n < m" by auto from lsm have "map_pmf (\y. y ! Suc n) (bv l) = map_pmf (\x. x!n) (bind_pmf (bv m) (\t. (return_pmf t)))" by (auto simp: map_bind_pmf) also have "\ = map_pmf (\x. x!n) (bv m)" by (auto simp: bind_return_pmf') also have "\ = bernoulli_pmf (5 / 10)" by (auto simp add: Suc(1)[of m, OF nltm]) finally show ?case . qed lemma pmf_2elemlist: "pmf (bv (Suc 0)) ([x]) = pmf (bv 0) [] * pmf (bernoulli_pmf (5 / 10)) x" unfolding bv.simps(2)[where n=0] pmf_bind pmf_return apply (subst integral_measure_pmf[where A="{[]}"]) apply (auto) by (cases x) auto lemma pmf_moreelemlist: "pmf (bv (Suc n)) (x#xs) = pmf (bv n) xs * pmf (bernoulli_pmf (5 / 10)) x" unfolding bv.simps(2) pmf_bind pmf_return apply (subst integral_measure_pmf[where A="{xs}"]) apply auto apply (cases x) apply(auto) apply (meson indicator_simps(2) list.inject singletonD) apply (meson indicator_simps(2) list.inject singletonD) apply (cases x) by(auto) lemma list_pmf: "length xs = n \ pmf (bv n) xs = (1 / 2)^n" proof(induct n arbitrary: xs) case 0 then have "xs = []" by auto then show "pmf (bv 0) xs = (1 / 2) ^ 0" by(auto) next case (Suc n xs) then obtain a as where split: "xs = a#as" by (metis Suc_length_conv) have "length as = n" using Suc(2) split by auto with Suc(1) have 1: "pmf (bv n) as = (1 / 2) ^ n" by auto from split pmf_moreelemlist[where n=n and x=a and xs=as] have "pmf (bv (Suc n)) xs = pmf (bv n) as * pmf (bernoulli_pmf (5 / 10)) a" by auto then have "pmf (bv (Suc n)) xs = (1 / 2) ^ n * 1 / 2" using 1 by auto then show "pmf (bv (Suc n)) xs = (1 / 2) ^ Suc n" by auto qed lemma bv_0_notlen: "pmf (bv n) xs = 0 \ length xs \ n " by(auto simp: list_pmf) lemma "length xs > n \ pmf (bv n) xs = 0" proof (induct n arbitrary: xs) case (Suc n xs) then obtain a as where split: "xs = a#as" by (metis Suc_length_conv Suc_lessE) have "length as > n" using Suc(2) split by auto with Suc(1) have 1: "pmf (bv n) as = 0" by auto from split pmf_moreelemlist[where n=n and x=a and xs=as] have "pmf (bv (Suc n)) xs = pmf (bv n) as * pmf (bernoulli_pmf (5 / 10)) a" by auto then have "pmf (bv (Suc n)) xs = 0 * 1 / 2" using 1 by auto then show "pmf (bv (Suc n)) xs = 0" by auto qed simp lemma map_hd_list_pmf: "map_pmf hd (bv (Suc n)) = bernoulli_pmf (5 / 10)" by (simp add: map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf') lemma map_tl_list_pmf: "map_pmf tl (bv (Suc n)) = bv n" by (simp add: map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf' ) subsection "function \flip\" fun flip :: "nat \ bool list \ bool list" where "flip _ [] = []" | "flip 0 (x#xs) = (\x)#xs" | "flip (Suc n) (x#xs) = x#(flip n xs)" lemma flip_length[simp]: "length (flip i xs) = length xs" apply(induct xs arbitrary: i) apply(simp) apply(case_tac i) by(simp_all) lemma flip_out_of_bounds: "y \ length X \ flip y X = X" apply(induct X arbitrary: y) proof - case (Cons X Xs) hence "y > 0" by auto with Cons obtain y' where y1: "y = Suc y'" and y2: "y' \ length Xs" by (metis Suc_pred' length_Cons not_less_eq_eq) then have "flip y (X # Xs) = X#(flip y' Xs)" by auto moreover from Cons y2 have "flip y' Xs = Xs" by auto ultimately show ?case by auto qed simp lemma flip_other: "y < length X \ z < length X \ z \ y \ flip z X ! y = X ! y" apply(induct y arbitrary: X z) apply(simp) apply (metis flip.elims neq0_conv nth_Cons_0) proof (case_tac z, goal_cases) case (1 y X z) then obtain a as where "X=a#as" using length_greater_0_conv by (metis (full_types) flip.elims) with 1(5) show ?case by(simp) next case (2 y X z z') from 2 have 3: "z' \ y" by auto from 2(2) have "length X > 0" by auto then obtain a as where aas: "X = a#as" by (metis (full_types) flip.elims length_greater_0_conv) then have a: "flip (Suc z') X ! Suc y = flip z' as ! y" and b : "(X ! Suc y) = (as ! y)" by auto from 2(2) aas have 1: "y < length as" by auto from 2(3,5) aas have f2: "z' < length as" by auto note c=2(1)[OF 1 f2 3] have "flip z X ! Suc y = flip (Suc z') X ! Suc y" using 2 by auto also have "\ = flip z' as ! y" by (rule a) also have "\ = as ! y" by (rule c) also have "\ = (X ! Suc y)" by (rule b[symmetric]) finally show "flip z X ! Suc y = (X ! Suc y)" . qed lemma flip_itself: "y < length X \ flip y X ! y = (\ X ! y)" apply(induct y arbitrary: X) apply(simp) apply (metis flip.elims nth_Cons_0 old.nat.distinct(2)) proof - fix y fix X::"bool list" assume iH: "(\X. y < length X \ flip y X ! y = (\ X ! y))" assume len: "Suc y < length X" from len have "y < length X" by auto from len have "length X > 0" by auto then obtain z zs where zzs: "X = z#zs" by (metis (full_types) flip.elims length_greater_0_conv) then have a: "flip (Suc y) X ! Suc y = flip y zs ! y" and b : "(\ X ! Suc y) = (\ zs ! y)" by auto from len zzs have "y < length zs" by auto note c=iH[OF this] from a b c show "flip (Suc y) X ! Suc y = (\ X ! Suc y)" by auto qed lemma flip_twice: "flip i (flip i b) = b" proof (cases "i < length b") case True then have A: "i < length (flip i b)" by simp show ?thesis apply(simp add: list_eq_iff_nth_eq) apply(clarify) proof (goal_cases) case (1 j) then show ?case apply(cases "i=j") using flip_itself[OF A] flip_itself[OF True] apply(simp) using flip_other True 1 by auto qed qed (simp add: flip_out_of_bounds) lemma flipidiflip: "y < length X \ e < length X \ flip e X ! y = (if e=y then ~ (X ! y) else X ! y)" apply(cases "e=y") apply(simp add: flip_itself) by(simp add: flip_other) lemma bernoulli_Not: "map_pmf Not (bernoulli_pmf (1 / 2)) = (bernoulli_pmf (1 / 2))" apply(rule pmf_eqI) proof (case_tac i, goal_cases) case (1 i) then have "pmf (map_pmf Not (bernoulli_pmf (1 / 2))) i = pmf (map_pmf Not (bernoulli_pmf (1 / 2))) (Not False)" by auto also have "\ = pmf (bernoulli_pmf (1 / 2)) False" apply (rule pmf_map_inj') apply(rule injI) by auto also have "\ = pmf (bernoulli_pmf (1 / 2)) i" by auto finally show ?case . next case (2 i) then have "pmf (map_pmf Not (bernoulli_pmf (1 / 2))) i = pmf (map_pmf Not (bernoulli_pmf (1 / 2))) (Not True)" by auto also have "\ = pmf (bernoulli_pmf (1 / 2)) True" apply (rule pmf_map_inj') apply(rule injI) by auto also have "\ = pmf (bernoulli_pmf (1 / 2)) i" by auto finally show ?case . qed lemma inv_flip_bv: "map_pmf (flip i) (bv n) = (bv n)" proof(induct n arbitrary: i) case (Suc n i) note iH=this have "bind_pmf (bv n) (\x. bind_pmf (bernoulli_pmf (1 / 2)) (\xa. map_pmf (flip i) (return_pmf (xa # x)))) = bind_pmf (bernoulli_pmf (1 / 2)) (\xa .bind_pmf (bv n) (\x. map_pmf (flip i) (return_pmf (xa # x))))" by(rule bind_commute_pmf) also have "\ = bind_pmf (bernoulli_pmf (1 / 2)) (\xa . bind_pmf (bv n) (\x. return_pmf (xa # x)))" proof (cases i) case 0 then have "bind_pmf (bernoulli_pmf (1 / 2)) (\xa. bind_pmf (bv n) (\x. map_pmf (flip i) (return_pmf (xa # x)))) = bind_pmf (bernoulli_pmf (1 / 2)) (\xa. bind_pmf (bv n) (\x. return_pmf ((\ xa) # x)))" by auto also have "\ = bind_pmf (bv n) (\x. bind_pmf (bernoulli_pmf (1 / 2)) (\xa. return_pmf ((\ xa) # x)))" by(rule bind_commute_pmf) also have "\ = bind_pmf (bv n) (\x. bind_pmf (map_pmf Not (bernoulli_pmf (1 / 2))) (\xa. return_pmf (xa # x)))" by(auto simp add: bind_map_pmf) also have "\ = bind_pmf (bv n) (\x. bind_pmf (bernoulli_pmf (1 / 2)) (\xa. return_pmf (xa # x)))" by (simp only: bernoulli_Not) also have "\ = bind_pmf (bernoulli_pmf (1 / 2)) (\xa. bind_pmf (bv n) (\x. return_pmf (xa # x)))" by(rule bind_commute_pmf) finally show ?thesis . next case (Suc i') have "bind_pmf (bernoulli_pmf (1 / 2)) (\xa. bind_pmf (bv n) (\x. map_pmf (flip i) (return_pmf (xa # x)))) = bind_pmf (bernoulli_pmf (1 / 2)) (\xa. bind_pmf (bv n) (\x. return_pmf (xa # flip i' x)))" unfolding Suc by(simp) also have "\ = bind_pmf (bernoulli_pmf (1 / 2)) (\xa. bind_pmf (map_pmf (flip i') (bv n)) (\x. return_pmf (xa # x)))" by(auto simp add: bind_map_pmf) also have "\ = bind_pmf (bernoulli_pmf (1 / 2)) (\xa. bind_pmf (bv n) (\x. return_pmf (xa # x)))" using iH[of "i'"] by simp finally show ?thesis . qed also have "\ = bind_pmf (bv n) (\x. bind_pmf (bernoulli_pmf (1 / 2)) (\xa. return_pmf (xa # x)))" by(rule bind_commute_pmf) finally show ?case by(simp add: map_pmf_def bind_assoc_pmf) qed simp subsection "Example for pmf" definition "twocoins = do { x \ (bernoulli_pmf 0.4); y \ (bernoulli_pmf 0.5); return_pmf (x \ y) }" lemma experiment0_7: "pmf twocoins True = 0.7" unfolding twocoins_def unfolding pmf_bind pmf_return apply (subst integral_measure_pmf[where A="{True, False}"]) by auto subsection "Sum Distribution" definition "Sum_pmf p Da Db = (bernoulli_pmf p) \ (%b. if b then map_pmf Inl Da else map_pmf Inr Db )" lemma b0: "bernoulli_pmf 0 = return_pmf False" apply(rule pmf_eqI) apply(case_tac i) by(simp_all) lemma b1: "bernoulli_pmf 1 = return_pmf True" apply(rule pmf_eqI) apply(case_tac i) by(simp_all) lemma Sum_pmf_0: "Sum_pmf 0 Da Db = map_pmf Inr Db" unfolding Sum_pmf_def apply(rule pmf_eqI) by(simp add: b0 bind_return_pmf) lemma Sum_pmf_1: "Sum_pmf 1 Da Db = map_pmf Inl Da" unfolding Sum_pmf_def apply(rule pmf_eqI) by(simp add: b1 bind_return_pmf) definition "Proj1_pmf D = map_pmf (%a. case a of Inl e \ e) (cond_pmf D {f. (\e. Inl e = f)})" lemma A: "(case_sum (\e. e) (\a. undefined)) (Inl e) = e" by(simp) lemma B: "inj (case_sum (\e. e) (\a. undefined))" oops lemma none: "p >0 \ p < 1 \ (set_pmf (bernoulli_pmf p \ (\b. if b then map_pmf Inl Da else map_pmf Inr Db)) \ {f. (\e. Inl e = f)}) \ {}" apply(simp add: UNIV_bool) using set_pmf_not_empty by fast lemma none2: "p >0 \ p < 1 \ (set_pmf (bernoulli_pmf p \ (\b. if b then map_pmf Inl Da else map_pmf Inr Db)) \ {f. (\e. Inr e = f)}) \ {}" apply(simp add: UNIV_bool) using set_pmf_not_empty by fast lemma C: "set_pmf (Proj1_pmf (Sum_pmf 0.5 Da Db)) = set_pmf Da" proof - show ?thesis unfolding Sum_pmf_def Proj1_pmf_def - applysimp + apply simp using none[of "0.5" Da Db] apply(simp add: set_cond_pmf UNIV_bool) by force qed thm integral_measure_pmf thm pmf_cond pmf_cond[OF none] lemma proj1_pmf: assumes "p>0" "p<1" shows "Proj1_pmf (Sum_pmf p Da Db) = Da" proof - have kl: "\e. pmf (map_pmf Inr Db) (Inl e) = 0" apply(simp only: pmf_eq_0_set_pmf) apply(simp) by blast have ll: "measure_pmf.prob (bernoulli_pmf p \ (\b. if b then map_pmf Inl Da else map_pmf Inr Db)) {f. \e. Inl e = f} = p" using assms apply(simp add: integral_pmf[symmetric] pmf_bind) apply(subst Bochner_Integration.integral_add) using integrable_pmf apply fast using integrable_pmf apply fast by(simp add: integral_pmf) have E: "(cond_pmf (bernoulli_pmf p \ (\b. if b then map_pmf Inl Da else map_pmf Inr Db)) {f. \e. Inl e = f}) = map_pmf Inl Da" apply(rule pmf_eqI) apply(subst pmf_cond) using none[of p Da Db] assms apply (simp) using assms apply(auto) apply(subst pmf_bind) apply(simp add: kl ll ) apply(simp only: pmf_eq_0_set_pmf) by auto have ID: "case_sum (\e. e) (\a. undefined) \ Inl = id" by fastforce show ?thesis unfolding Sum_pmf_def Proj1_pmf_def apply(simp only: E) apply(simp add: pmf.map_comp ID) done qed definition "Proj2_pmf D = map_pmf (%a. case a of Inr e \ e) (cond_pmf D {f. (\e. Inr e = f)})" lemma proj2_pmf: assumes "p>0" "p<1" shows "Proj2_pmf (Sum_pmf p Da Db) = Db" proof - have kl: "\e. pmf (map_pmf Inl Da) (Inr e) = 0" apply(simp only: pmf_eq_0_set_pmf) apply(simp) by blast have ll: "measure_pmf.prob (bernoulli_pmf p \ (\b. if b then map_pmf Inl Da else map_pmf Inr Db)) {f. \e. Inr e = f} = 1-p" using assms apply(simp add: integral_pmf[symmetric] pmf_bind) apply(subst Bochner_Integration.integral_add) using integrable_pmf apply fast using integrable_pmf apply fast by(simp add: integral_pmf) have E: "(cond_pmf (bernoulli_pmf p \ (\b. if b then map_pmf Inl Da else map_pmf Inr Db)) {f. \e. Inr e = f}) = map_pmf Inr Db" apply(rule pmf_eqI) apply(subst pmf_cond) using none2[of p Da Db] assms apply (simp) using assms apply(auto) apply(subst pmf_bind) apply(simp add: kl ll ) apply(simp only: pmf_eq_0_set_pmf) by auto have ID: "case_sum (\e. undefined) (\a. a) \ Inr = id" by fastforce show ?thesis unfolding Sum_pmf_def Proj2_pmf_def apply(simp only: E) apply(simp add: pmf.map_comp ID) done qed definition "invSum invA invB D x i == invA (Proj1_pmf D) x i \ invB (Proj2_pmf D) x i" lemma invSum_split: "p>0 \ p<1 \ invA Da x i \ invB Db x i \ invSum invA invB (Sum_pmf p Da Db) x i" by(simp add: invSum_def proj1_pmf proj2_pmf) term "(%a. case a of Inl e \ Inl (fa e) | Inr e \ Inr (fb e))" definition "f_on2 fa fb = (%a. case a of Inl e \ map_pmf Inl (fa e) | Inr e \ map_pmf Inr (fb e))" term "bind_pmf" lemma Sum_bind_pmf: assumes a: "bind_pmf Da fa = Da'" and b: "bind_pmf Db fb = Db'" shows "bind_pmf (Sum_pmf p Da Db) (f_on2 fa fb) = Sum_pmf p Da' Db'" proof - { fix x have "(if x then map_pmf Inl Da else map_pmf Inr Db) \ case_sum (\e. map_pmf Inl (fa e)) (\e. map_pmf Inr (fb e)) = (if x then map_pmf Inl Da \ case_sum (\e. map_pmf Inl (fa e)) (\e. map_pmf Inr (fb e)) else map_pmf Inr Db \ case_sum (\e. map_pmf Inl (fa e)) (\e. map_pmf Inr (fb e)))" apply(simp) done also have "\ = (if x then map_pmf Inl (bind_pmf Da fa) else map_pmf Inr (bind_pmf Db fb))" by(auto simp add: map_pmf_def bind_assoc_pmf bind_return_pmf) also have "\ = (if x then map_pmf Inl Da' else map_pmf Inr Db')" using a b by simp finally have "(if x then map_pmf Inl Da else map_pmf Inr Db) \ case_sum (\e. map_pmf Inl (fa e)) (\e. map_pmf Inr (fb e)) = (if x then map_pmf Inl Da' else map_pmf Inr Db')" . } note gr=this show ?thesis unfolding Sum_pmf_def f_on2_def apply(rule pmf_eqI) apply(case_tac i) by(simp_all add: bind_return_pmf bind_assoc_pmf gr) qed definition "sum_map_pmf fa fb = (%a. case a of Inl e \ Inl (fa e) | Inr e \ Inr (fb e))" lemma Sum_map_pmf: assumes a: "map_pmf fa Da = Da'" and b: "map_pmf fb Db = Db'" shows "map_pmf (sum_map_pmf fa fb) (Sum_pmf p Da Db) = Sum_pmf p Da' Db'" proof - have "map_pmf (sum_map_pmf fa fb) (Sum_pmf p Da Db) = bind_pmf (Sum_pmf p Da Db) (f_on2 (\x. return_pmf (fa x)) (\x. return_pmf (fb x)))" using a b unfolding map_pmf_def sum_map_pmf_def f_on2_def by(auto simp add: bind_return_pmf sum.case_distrib) also have "\ = Sum_pmf p Da' Db'" using assms[unfolded map_pmf_def] by(rule Sum_bind_pmf ) finally show ?thesis . qed end diff --git a/thys/List_Update/TS.thy b/thys/List_Update/TS.thy --- a/thys/List_Update/TS.thy +++ b/thys/List_Update/TS.thy @@ -1,2682 +1,2682 @@ (* Title: Competitive Analysis of TS Author: Max Haslbeck *) section "TS: another 2-competitive Algorithm" theory TS imports OPT2 Phase_Partitioning Move_to_Front List_Factoring RExp_Var begin subsection "Definition of TS" definition TS_step_d where "TS_step_d s q = (( ( let li = index (snd s) q in (if li = length (snd s) then 0 \ \requested for first time\ else (let sincelast = take li (snd s) in (let S={x. x < q in (fst s) \ count_list sincelast x \ 1} in (if S={} then 0 else (index (fst s) q) - Min ( (index (fst s)) ` S))) ) ) ) ,[]), q#(snd s))" (* FIXME: generalizing regular expressions equivalence checking enables relaxing the type here to 'a::linord *) definition rTS :: "nat list \ (nat,nat list) alg_on" where "rTS h = ((\s. h), TS_step_d)" fun TSstep where "TSstep qs n (is,s) = ((qs!n)#is, step s (qs!n) (( let li = index is (qs!n) in (if li = length is then 0 \ \requested for first time\ else (let sincelast = take li is in (let S={x. x < (qs!n) in s \ count_list sincelast x \ 1} in (if S={} then 0 else (index s (qs!n)) - Min ( (index s) ` S))) ) ) ),[]))" lemma TSnopaid: "(snd (fst (snd (rTS initH) is q))) = []" unfolding rTS_def by(simp add: TS_step_d_def) abbreviation TSdet where "TSdet init initH qs n == config (rTS initH) init (take n qs)" lemma TSdet_Suc: "Suc n \ length qs \ TSdet init initH qs (Suc n) = Step (rTS initH) (TSdet init initH qs n) (qs!n)" by(simp add: take_Suc_conv_app_nth config_snoc) (* now do the proof with TSdet *) definition s_TS where "s_TS init initH qs n = fst (TSdet init initH qs n)" lemma sndTSdet: "n\length xs \ snd (TSdet init initH xs n) = rev (take n xs) @ initH" apply(induct n) apply(simp add: rTS_def) by(simp add: split_def TS_step_d_def take_Suc_conv_app_nth config'_snoc Step_def rTS_def) subsection "Behaviour of TS on lists of length 2" lemma fixes hs x y assumes "x\y" shows oneTS_step : "TS_step_d ([x, y], x#y#hs) y = ((1, []), y # x # y # hs)" and oneTS_stepyyy: "TS_step_d ([x, y], y#x#hs) y = ((Suc 0, []), y#y#x#hs)" and oneTS_stepx: "TS_step_d ([x, y], x#x#hs) y = ((0, []), y # x # x # hs)" and oneTS_stepy: "TS_step_d ([x, y], []) y = ((0, []), [y])" and oneTS_stepxy: "TS_step_d ([x, y], [x]) y = ((0, []), [y, x])" and oneTS_stepyy: "TS_step_d ([x, y], [y]) y = ((Suc 0, []), [y, y])" and oneTS_stepyx: "TS_step_d ([x, y], hs) x = ((0, []), x # hs)" using assms by(auto simp add: step_def mtf2_def swap_def TS_step_d_def before_in_def) lemmas oneTS_steps = oneTS_stepx oneTS_stepxy oneTS_stepyx oneTS_stepy oneTS_stepyy oneTS_stepyyy oneTS_step subsection "Analysis of the Phases" definition "TS_inv c x i \ (\hs. c = return_pmf ((if x=hd i then i else rev i),[x,x]@hs) ) \ c = return_pmf ((if x=hd i then i else rev i),[])" lemma TS_inv_sym: "a\b \ {a,b}={x,y} \ z\{x,y} \ TS_inv c z [a,b] = TS_inv c z [x,y]" unfolding TS_inv_def by auto abbreviation "TS_inv' s x i == TS_inv (return_pmf s) x i" lemma TS_inv'_det: "TS_inv' s x i = ((\hs. s = ((if x=hd i then i else rev i),[x,x]@hs) ) \ s = ((if x=hd i then i else rev i),[]))" unfolding TS_inv_def by auto lemma TS_inv'_det2: "TS_inv' (s,h) x i = (\hs. (s,h) = ((if x=hd i then i else rev i),[x,x]@hs) ) \ (s,h) = ((if x=hd i then i else rev i),[])" unfolding TS_inv_def by auto (* TS_A (x+1)yy \ Plus(Atom (x::nat)) One,(Atom y), (Atom y)] TS_B (x+1)yx(yx)*yy \ Plus(Atom x) One,(Atom y),(Atom x),Star(Times (Atom y)(Atom x)),(Atom y),(Atom y)] TS_C (x+1)yx(yx)*x \ Plus(Atom x) One,(Atom y),(Atom x),Star(Times (Atom y)(Atom x)),(Atom x)] TD_D xx \ seq[(Atom x),(Atom x)] *) subsubsection "(yx)*?" lemma TS_yx': assumes "x \ y" "qs \ lang (Star(Times (Atom y) (Atom x)))" "\hs. h=[x,y]@hs" shows "T_on' (rTS h0) ([x,y],h) (qs@r) = length qs + T_on' (rTS h0) ([x,y],((rev qs) @h)) r \ (\hs. ((rev qs) @h) = [x, y] @ hs) \ config' (rTS h0) ([x, y],h) qs = ([x,y],rev qs @ h)" proof - from assms have "qs \ star ({[y]} @@ {[x]})" by (simp) from this assms(3) show ?thesis proof (induct qs arbitrary: h rule: star_induct) case Nil then show ?case by(simp add: rTS_def) next case (append u v) then have uyx: "u = [y,x]" by auto from append obtain hs where a: "h = [x,y]@hs" by blast have "T_on' (rTS h0) ([x, y], (rev u @ h)) (v @ r) = length v + T_on' (rTS h0) ([x, y], rev v @ (rev u @ h)) r \ (\hs. rev v @ (rev u @ h) = [x, y] @ hs) \ config' (rTS h0) ([x, y], (rev u @ h)) v = ([x, y], rev v @ (rev u @ h))" apply(simp only: uyx) apply(rule append(3)) by simp then have yy: "T_on' (rTS h0) ([x, y], (rev u @ h)) (v @ r) = length v + T_on' (rTS h0) ([x, y], rev v @ (rev u @ h)) r" and history: "(\hs. rev v @ (rev u @ h) = [x, y] @ hs)" and state: "config' (rTS h0) ([x, y], (rev u @ h)) v = ([x, y], rev v @ (rev u @ h))" by auto have s0: "s_TS [x, y] h [y, x] 0 = [x,y]" unfolding s_TS_def by(simp) from assms(1) have hahah: " {xa. xa < y in [x, y] \ count_list [x] xa \ 1} = {x}" unfolding before_in_def by auto have "config' (rTS h0) ([x, y],h) u = ([x, y], x # y # x # y # hs)" apply(simp add: split_def rTS_def uyx a ) using assms(1) by(auto simp add: Step_def oneTS_steps step_def mtf2_def swap_def) then have s2: "config' (rTS h0) ([x, y],h) u = ([x, y], ((rev u) @ h))" unfolding a uyx by simp have "config' (rTS h0) ([x, y], h) (u @ v) = config' (rTS h0) (Partial_Cost_Model.config' (rTS h0) ([x, y], h) u) v" by (rule config'_append2) also have "\ = config' (rTS h0) ([x, y], ((rev u) @ h)) v" by(simp only: s2) also have "\ = ([x, y], rev (u @ v) @ h)" by (simp add: state) finally have alles: "config' (rTS h0) ([x, y], h) (u @ v) = ([x, y], rev (u @ v) @ h)" . have ta: "T_on' (rTS h0) ([x,y],h) u = 2" unfolding rTS_def uyx a apply(simp only: T_on'.simps(2)) using assms(1) apply(auto simp add: Step_def step_def mtf2_def swap_def oneTS_steps) by(simp add: t\<^sub>p_def) have "T_on' (rTS h0) ([x,y],h) ((u @ v) @ r) = T_on' (rTS h0) ([x,y],h) (u @ (v @ r))" by auto also have "\ = T_on' (rTS h0) ([x,y],h) u + T_on' (rTS h0) (config' (rTS h0) ([x, y],h) u) (v @ r)" by(rule T_on'_append) also have "\ = T_on' (rTS h0) ([x,y],h) u + T_on' (rTS h0) ([x, y],(rev u @ h)) (v @ r)" by(simp only: s2) also have "\ = T_on' (rTS h0) ([x,y],h) u + length v + T_on' (rTS h0) ([x, y], rev v @ (rev u @ h)) r" by(simp only: yy) also have "\ = 2 + length v + T_on' (rTS h0) ([x, y], rev v @ (rev u @ h)) r" by(simp only: ta) also have "\ = length (u @ v) + T_on' (rTS h0) ([x, y], rev v @ (rev u @ h)) r" using uyx by auto also have "\ = length (u @ v) + T_on' (rTS h0) ([x, y], (rev (u @ v) @ h)) r" by auto finally show ?case using history alles by simp qed qed subsubsection "?x" lemma TS_x': "T_on' (rTS h0) ([x,y],h) [x] = 0 \ config' (rTS h0) ([x, y],h) [x] = ([x,y], rev [x] @ h)" by(auto simp add: t\<^sub>p_def rTS_def TS_step_d_def Step_def step_def) subsubsection "?yy" lemma TS_yy': assumes "x \ y" "\hs. h = [x, y] @ hs" shows "T_on' (rTS h0) ([x,y],h) [y, y] = 1" "config' (rTS h0) ([x, y],h) [y,y] = ([y,x],rev [y,y] @ h)" proof - from assms obtain hs where a: "h = [x,y]@hs" by blast from a show "T_on' (rTS h0) ([x,y],h) [y, y] = 1" unfolding rTS_def using assms(1) apply(auto simp add: oneTS_steps Step_def step_def mtf2_def swap_def) by(simp add: t\<^sub>p_def) show "config' (rTS h0) ([x, y],h) [y,y] = ([y,x],rev [y,y] @ h)" unfolding rTS_def a using assms(1) by(simp add: Step_def oneTS_steps step_def mtf2_def swap_def) qed subsubsection "yx(yx)*?" lemma TS_yxyx': assumes [simp]: "x \ y" and "qs \ lang (seq[Times (Atom y) (Atom x), Star(Times (Atom y) (Atom x))])" "(\hs. h=[x,x]@hs) \ index h y = length h" shows "T_on' (rTS h0) ([x,y],h) (qs@r) = length qs - 1 + T_on' (rTS h0) ([x,y],rev qs @ h) r \ (\hs. (rev qs @ h) = [x, y] @ hs) \ config' (rTS h0) ([x, y],h) qs = ([x,y], rev qs @ h)" proof - obtain u v where uu: "u \ lang (Times (Atom y) (Atom x))" and vv: "v \ lang (seq[ Star(Times (Atom y) (Atom x))])" and qsuv: "qs = u @ v" using assms(2) by (auto simp: conc_def) from uu have uyx: "u = [y,x]" by(auto) from qsuv uyx have vqs: "length v = length qs - 2" by auto from qsuv uyx have vqs2: "length v + 1 = length qs - 1" by auto have firststep: "TS_step_d ([x, y], h) y = ((0, []), y # h)" proof (cases "index h y = length h") case True then show ?thesis unfolding TS_step_d_def by(simp) next case False with assms(3) obtain hs where a: "h = [x,x]@hs" by auto then show ?thesis by(simp add: oneTS_steps) qed have s2: "config' (rTS h0) ([x,y],h) u = ([x, y], x # y # h)" - unfolding rTS_def uyx applysimp + unfolding rTS_def uyx apply simp unfolding Step_def by(simp add: firststep step_def oneTS_steps) have ta: "T_on' (rTS h0) ([x,y],h) u = 1" unfolding rTS_def uyx apply(simp) apply(simp add: firststep) unfolding Step_def using assms(1) by (simp add: firststep step_def oneTS_steps t\<^sub>p_def) have ttt: "T_on' (rTS h0) ([x,y],rev u @ h) (v@r) = length v + T_on' (rTS h0) ([x,y],((rev v) @(rev u @ h))) r \ (\hs. ((rev v) @(rev u @ h)) = [x, y] @ hs) \ config' (rTS h0) ([x, y],(rev u @ h)) v = ([x,y],rev v @ (rev u @ h))" apply(rule TS_yx') apply(fact) using vv apply(simp) using uyx by(simp) then have tat: "T_on' (rTS h0) ([x,y], x # y # h) (v@r) = length v + T_on' (rTS h0) ([x,y],rev qs @ h) r" and history: "(\hs. (rev qs @ h) = [x, y] @ hs)" and state: "config' (rTS h0) ([x, y], x # y # h) v = ([x,y],rev qs @ h)" using qsuv uyx by auto have "config' (rTS h0) ([x, y], h) qs = config' (rTS h0) (config' (rTS h0) ([x, y], h) u) v" unfolding qsuv by (rule config'_append2) also have "\ = ([x, y], rev qs @ h)" by(simp add: s2 state) finally have his: "config' (rTS h0) ([x, y], h) qs = ([x, y], rev qs @ h)" . have "T_on' (rTS h0) ([x,y],h) (qs@r) = T_on' (rTS h0) ([x,y],h) (u @ v @ r)" using qsuv by auto also have "\ = T_on' (rTS h0) ([x,y],h) u + T_on' (rTS h0) (config' (rTS h0) ([x,y],h) u) (v @ r)" by(rule T_on'_append) also have "\ = T_on' (rTS h0) ([x,y],h) u + T_on' (rTS h0) ([x, y], x # y # h) (v @ r)" by(simp only: s2) also have "\ = T_on' (rTS h0) ([x,y],h) u + length v + T_on' (rTS h0) ([x,y],rev qs @ h) r" by (simp only: tat) also have "\ = 1 + length v + T_on' (rTS h0) ([x,y],rev qs @ h) r" by(simp only: ta) also have "\ = length qs - 1 + T_on' (rTS h0) ([x,y],rev qs @ h) r" using vqs2 by auto finally show ?thesis apply(safe) using history apply(simp) using his by auto qed lemma TS_xr': assumes "x \ y" "qs \ lang (Plus (Atom x) One)" "h = [] \ (\hs. h = [x, x] @ hs) " shows "T_on' (rTS h0) ([x,y],h) (qs@r) = T_on' (rTS h0) ([x,y],rev qs@h) r" "((\hs. (rev qs @ h) = [x, x] @ hs) \ (rev qs @ h) = [x] \ (rev qs @ h)=[]) " "config' (rTS h0) ([x,y],h) (qs@r) = config' (rTS h0) ([x,y],rev qs @ h) r" using assms by (auto simp add: T_on'_append Step_def rTS_def TS_step_d_def step_def t\<^sub>p_def) subsubsection "(x+1)yx(yx)*yy" lemma ts_b': assumes "x \ y" "v \ lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom y, Atom y])" "(\hs. h = [x, x] @ hs) \ h = [x] \ h = []" shows "T_on' (rTS h0) ([x, y], h) v = (length v - 2) \ (\hs. (rev v @ h) = [y,y]@hs) \ config' (rTS h0) ([x,y], h) v = ([y,x], rev v @ h)" proof - from assms have lenvmod: "length v mod 2 = 0" apply(simp) proof - assume "v \ ({[y]} @@ {[x]}) @@ star ({[y]} @@ {[x]}) @@ {[y]} @@ {[y]}" then obtain p q r where pqr: "v=p@q@r" and "p\({[y]} @@ {[x]})" and q: "q \ star ({[y]} @@ {[x]})" and "r \ {[y]} @@ {[y]}" by (metis concE) then have "p = [y,x]" "r=[y,y]" by auto with pqr have a: "length v = 4+length q" by auto from q have b: "length q mod 2 = 0" apply(induct q rule: star_induct) by (auto) from a b show ?thesis by auto qed with assms(1,3) have fall: "(\hs. h = [x, x] @ hs) \ index h y = length h" by(auto) from assms(2) have "v \ lang (seq[Times (Atom y) (Atom x), Star(Times (Atom y) (Atom x))]) @@ lang (seq[Atom y, Atom y])" by (auto simp: conc_def) then obtain a b where aa: "a \ lang (seq[Times (Atom y) (Atom x), Star(Times (Atom y) (Atom x))])" and "b \ lang (seq[Atom y, Atom y])" and vab: "v = a @ b" by(erule concE) then have bb: "b=[y,y]" by auto from aa have lena: "length a > 0" by auto from TS_yxyx'[OF assms(1) aa fall] have stars: "T_on' (rTS h0) ([x, y], h) (a @ b) = length a - 1 + T_on' (rTS h0) ([x, y], rev a @ h) b" and history: "(\hs. rev a @ h = [x, y] @ hs)" and state: "config' (rTS h0) ([x, y], h) a = ([x,y],rev a @ h)" by auto (* "T_on' (rTS h0) ([x,y],h) [y, y] = 1" "config' (rTS h0) ([x, y],h) [y,y] = ([y,x],rev [y,y] @ h)" *) have suffix: "T_on' (rTS h0) ([x, y], rev a @ h) b = 1" and jajajaj: "config' (rTS h0) ([x, y],rev a @ h) b = ([y,x],rev b @ rev a @ h)" unfolding bb using TS_yy' history assms(1) by auto from stars suffix have "T_on' (rTS h0) ([x, y], h) (a @ b) = length a" using lena by auto then have whatineed: "T_on' (rTS h0) ([x, y], h) v = (length v - 2)" using vab bb by auto have grgr:"config' (rTS h0) ([x, y], h) v = ([y, x], rev v @ h)" unfolding vab apply(simp only: config'_append2 state jajajaj) by simp from history obtain hs' where "rev a @ h = [x, y] @ hs'" by auto then obtain hs2 where reva: "rev a @ h = x # hs2" by auto show ?thesis using whatineed grgr by(auto simp add: reva vab bb) qed lemma TS_b'1: assumes "x \ y" "h = [] \ (\hs. h = [x, x] @ hs)" "qs \ lang (seq [Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom y, Atom y])" shows "T_on' (rTS h0) ([x, y], h) qs = (length qs - 2) \ TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x,y]" proof - have f: "qs \ lang (seq [Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom y, Atom y])" using assms(3) by(simp add: conc_assoc) from ts_b'[OF assms(1) f] assms(2) have T_star: "T_on' (rTS h0) ([x, y], h) qs = length qs - 2" and inv1: "config' (rTS h0) ([x, y], h) qs = ([y, x], rev qs @ h)" and inv2: "(\hs. rev qs @ h = [y, y] @ hs)" by auto from T_star have TS: "T_on' (rTS h0) ([x, y], h) qs = (length qs - 2)" by metis have lqs: "last qs = y" using assms(3) by force from inv1 have inv: "TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x, y]" apply(simp add: lqs) apply(subst TS_inv'_det) using assms(2) inv2 by(simp) show ?thesis unfolding TS apply(safe) by(fact inv) qed lemma TS_b1'': assumes "x \ y" "{x, y} = {x0, y0}" "TS_inv s x [x0, y0]" "set qs \ {x, y}" "qs \ lang (seq [Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom y, Atom y])" shows "TS_inv (config'_rand (embed (rTS h0)) s qs) (last qs) [x0, y0] \ T_on_rand' (embed (rTS h0)) s qs = (length qs - 2)" proof - from assms(1,2) have kas: "(x0=x \ y0=y) \ (y0=x \ x0=y)" by(auto) then obtain h where S: "s = return_pmf ([x,y],h)" and h: "h = [] \ (\hs. h = [x, x] @ hs)" apply(rule disjE) using assms(1,3) unfolding TS_inv_def by(auto) have l: "qs \ []" using assms by auto { fix x y qs h0 fix h:: "nat list" assume A: "x \ y" and B: "qs \ lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom y, Atom y])" and C: "h = [] \ (\hs. h = [x, x] @ hs)" then have C': "(\hs. h = [x, x] @ hs) \ h = [x] \ h = []" by blast from B have lqs: "last qs = y" using assms(5) by(auto simp add: conc_def) have "TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) qs) (last qs) [x, y] \ T_on_rand' (embed (rTS h0)) (return_pmf ([x, y], h)) qs = length qs - 2" apply(simp only: T_on'_embed[symmetric] config'_embed) using ts_b'[OF A B C'] A lqs unfolding TS_inv'_det by auto } note b1=this show ?thesis unfolding S using kas apply(rule disjE) apply(simp only:) apply(rule b1) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) apply(simp only:) apply(subst TS_inv_sym[of y x x y]) using assms(1) apply(simp) apply(blast) defer apply(rule b1) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) using last_in_set l assms(4) by blast qed lemma ts_b2': assumes "x \ y" "qs \ lang (seq[Atom x, Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom y, Atom y])" "(\hs. h = [x, x] @ hs) \ h = []" shows "T_on' (rTS h0) ([x, y], h) qs = (length qs - 3) \ config' (rTS h0) ([x,y], h) qs = ([y,x],rev qs@h) \ (\hs. (rev qs @ h) = [y,y]@hs)" proof - from assms(2) obtain v where qs: "qs = [x]@v" and V: "v\lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom y, Atom y])" by(auto simp add: conc_assoc) from assms(3) have 3: "(\hs. x#h = [x, x] @ hs) \ x#h = [x] \ x#h = []" by auto from ts_b'[OF assms(1) V 3] have T: "T_on' (rTS h0) ([x, y], x#h) v = length v - 2" and C: "config' (rTS h0) ([x, y], x#h) v = ([y, x], rev v @ x#h)" and H: "(\hs. rev v @ x#h = [y, y] @ hs)" by auto have t: "t\<^sub>p [x, y] x (fst (snd (rTS h0) ([x, y], h) x)) = 0" by (simp add: step_def rTS_def TS_step_d_def t\<^sub>p_def) have c: "Partial_Cost_Model.Step (rTS h0) ([x, y], h) x = ([x,y], x#h)" by (simp add: Step_def rTS_def TS_step_d_def step_def) show ?thesis unfolding qs apply(safe) apply(simp add: T_on'_append T c t) apply(simp add: config'_rand_append C c) using H by simp qed lemma TS_b2'': assumes "x \ y" "{x, y} = {x0, y0}" "TS_inv s x [x0, y0]" "set qs \ {x, y}" "qs \ lang (seq [Atom x, Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom y, Atom y])" shows "TS_inv (config'_rand (embed (rTS h0)) s qs) (last qs) [x0, y0] \ T_on_rand' (embed (rTS h0)) s qs = (length qs - 3)" proof - from assms(1,2) have kas: "(x0=x \ y0=y) \ (y0=x \ x0=y)" by(auto) then obtain h where S: "s = return_pmf ([x,y],h)" and h: "h = [] \ (\hs. h = [x, x] @ hs)" apply(rule disjE) using assms(1,3) unfolding TS_inv_def by(auto) have l: "qs \ []" using assms by auto { fix x y qs h0 fix h:: "nat list" assume A: "x \ y" and B: "qs \ lang (seq[Atom x, Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom y, Atom y])" and C: "h = [] \ (\hs. h = [x, x] @ hs)" from B have lqs: "last qs = y" using assms(5) by(auto simp add: conc_def) from C have C': "(\hs. h = [x, x] @ hs) \ h = []" by blast have "TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) qs) (last qs) [x, y] \ T_on_rand' (embed (rTS h0)) (return_pmf ([x, y], h)) qs = length qs - 3" apply(simp only: T_on'_embed[symmetric] config'_embed) using ts_b2'[OF A B C'] A lqs unfolding TS_inv'_det by auto } note b2=this show ?thesis unfolding S using kas apply(rule disjE) apply(simp only:) apply(rule b2) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) apply(simp only:) apply(subst TS_inv_sym[of y x x y]) using assms(1) apply(simp) apply(blast) defer apply(rule b2) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) using last_in_set l assms(4) by blast qed lemma TS_b': assumes "x \ y" "h = [] \ (\hs. h = [x, x] @ hs)" "qs \ lang (seq [Plus (Atom x) rexp.One, Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom y, Atom y])" shows "T_on' (rTS h0) ([x, y], h) qs \ 2 * T\<^sub>p [x, y] qs (OPT2 qs [x, y]) \ TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x,y]" proof - obtain u v where uu: "u \ lang (Plus (Atom x) One)" and vv: "v \ lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom y, Atom y])" and qsuv: "qs = u @ v" using assms(3) by (auto simp: conc_def) from TS_xr'[OF assms(1) uu assms(2)] have T_pre: "T_on' (rTS h0) ([x, y], h) (u @ v) = T_on' (rTS h0) ([x, y], rev u @ h) v" and fall': "(\hs. (rev u @ h) = [x, x] @ hs) \ (rev u @ h) = [x] \ (rev u @ h)=[]" and conf: "config' (rTS h0) ([x,y],h) (u@v) = config' (rTS h0) ([x,y],rev u @ h) v" by auto with assms uu have fall: "(\hs. (rev u @ h) = [x, x] @ hs) \ index (rev u @ h) y = length (rev u @ h)" by(auto) from ts_b'[OF assms(1) vv fall'] have T_star: "T_on' (rTS h0) ([x, y], rev u @ h) v = length v - 2" and inv1: "config' (rTS h0) ([x, y], rev u @ h) v = ([y, x], rev v @ rev u @ h)" and inv2: "(\hs. rev v @ rev u @ h = [y, y] @ hs)" by auto from T_pre T_star qsuv have TS: "T_on' (rTS h0) ([x, y], h) qs = (length v - 2)" by metis (* OPT *) from uu have uuu: "u=[] \ u=[x]" by auto from vv have vvv: "v \ lang (seq [Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom y, Atom y])" by(auto simp: conc_def) have OPT: "T\<^sub>p [x,y] qs (OPT2 qs [x,y]) = (length v) div 2" apply(rule OPT2_B) by(fact)+ have lqs: "last qs = y" using assms(3) by force have "config' (rTS h0) ([x, y], h) qs = ([y, x], rev qs @ h)" unfolding qsuv conf inv1 by simp then have inv: "TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x, y]" apply(simp add: lqs) apply(subst TS_inv'_det) using assms(2) inv2 qsuv by(simp) show ?thesis unfolding TS OPT apply(safe) apply(simp) by(fact inv) qed subsubsection "(x+1)yy" lemma ts_a': assumes "x \ y" "qs \ lang (seq [Plus (Atom x) One, Atom y, Atom y])" "h = [] \ (\hs. h = [x, x] @ hs)" shows "TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x,y] \ T_on' (rTS h0) ([x, y], h) qs = 2" proof - obtain u v where uu: "u \ lang (Plus (Atom x) One)" and vv: "v \ lang (seq[Atom y, Atom y])" and qsuv: "qs = u @ v" using assms(2) by (auto simp: conc_def) from vv have vv2: "v = [y,y]" by auto from uu have TS_prefix: " T_on' (rTS h0) ([x, y], h) u = 0" using assms(1) by(auto simp add: rTS_def oneTS_steps t\<^sub>p_def) have h_split: "rev u @ h = [] \ rev u @ h = [x] \ (\ hs. rev u @ h = [x,x]@hs)" using assms(3) uu by(auto) then have e: "T_on' (rTS h0) ([x,y],rev u @ h) [y,y] = 2" using assms(1) apply(auto simp add: rTS_def oneTS_steps Step_def step_def t\<^sub>p_def) done have conf: "config' (rTS h0) ([x, y], h) u = ([x,y], rev u @ h)" using uu by(auto simp add: Step_def rTS_def TS_step_d_def step_def) have "T_on' (rTS h0) ([x, y], h) qs = T_on' (rTS h0) ([x, y], h) (u @ v)" using qsuv by auto also have "\ =T_on' (rTS h0) ([x, y], h) u + T_on' (rTS h0) (config' (rTS h0) ([x, y], h) u) v" by(rule T_on'_append) also have "\ = T_on' (rTS h0) ([x, y], h) u + T_on' (rTS h0) ([x,y],rev u @ h) [y,y]" by(simp add: conf vv2) also have "\ = T_on' (rTS h0) ([x, y], h) u + 2" by (simp only: e) also have "\ = 2" by (simp add: TS_prefix) finally have TS: "T_on' (rTS h0) ([x, y], h) qs= 2" . (* dannach *) have lqs: "last qs = y" using assms(2) by force from assms(1) have "config' (rTS h0) ([x, y], h) qs = ([y,x], rev qs @ h)" unfolding qsuv apply(simp only: config'_append2 conf vv2) using h_split apply(auto simp add: Step_def rTS_def oneTS_steps step_def) by(simp_all add: mtf2_def swap_def) with assms(1) have "TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x,y]" apply(subst TS_inv'_det) by(simp add: qsuv vv2 lqs) show ?thesis unfolding TS apply(auto) by fact qed lemma TS_a': assumes "x \ y" "h = [] \ (\hs. h = [x, x] @ hs)" and "qs \ lang (seq [Plus (Atom x) rexp.One, Atom y, Atom y])" shows "T_on' (rTS h0) ([x, y], h) qs \ 2 * T\<^sub>p [x, y] qs (OPT2 qs [x, y]) \ TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x, y] \ T_on' (rTS h0) ([x, y], h) qs = 2" proof - have OPT: "T\<^sub>p [x,y] qs (OPT2 qs [x,y]) = 1" using OPT2_A[OF assms(1,3)] by auto show ?thesis using OPT ts_a'[OF assms(1,3,2)] by auto qed lemma TS_a'': assumes "x \ y" "{x, y} = {x0, y0}" "TS_inv s x [x0, y0]" "set qs \ {x, y}" "qs \ lang (seq [Plus (Atom x) One, Atom y, Atom y])" shows "TS_inv (config'_rand (embed (rTS h0)) s qs) (last qs) [x0, y0] \ T\<^sub>p_on_rand' (embed (rTS h0)) s qs = 2" proof - from assms(1,2) have kas: "(x0=x \ y0=y) \ (y0=x \ x0=y)" by(auto) then obtain h where S: "s = return_pmf ([x,y],h)" and h: "h = [] \ (\hs. h = [x, x] @ hs)" apply(rule disjE) using assms(1,3) unfolding TS_inv_def by(auto) have l: "qs \ []" using assms by auto { fix x y qs h0 fix h:: "nat list" assume A: "x \ y" "qs \ lang (seq [question (Atom x), Atom y, Atom y])" "h = [] \ (\hs. h = [x, x] @ hs)" have "TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) qs) (last qs) [x, y] \ T_on_rand' (embed (rTS h0)) (return_pmf ([x, y], h)) qs = 2" apply(simp only: T_on'_embed[symmetric] config'_embed) using ts_a'[OF A] by auto } note b=this show ?thesis unfolding S using kas apply(rule disjE) apply(simp only:) apply(rule b) using assms apply(simp) using assms apply(simp) using h apply(simp) apply(simp only:) apply(subst TS_inv_sym[of y x x y]) using assms(1) apply(simp) apply(blast) defer apply(rule b) using assms apply(simp) using assms apply(simp) using h apply(simp) using last_in_set l assms(4) by blast qed subsubsection "x+yx(yx)*x" lemma ts_c': assumes "x \ y" "v \ lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom x])" "(\hs. h = [x, x] @ hs) \ h = [x] \ h = []" shows "T_on' (rTS h0) ([x, y], h) v = (length v - 2) \ config' (rTS h0) ([x,y], h) v = ([x,y],rev v@h) \ (\hs. (rev v @ h) = [x,x]@hs)" proof - from assms have lenvmod: "length v mod 2 = 1" apply(simp) proof - assume "v \ ({[y]} @@ {[x]}) @@ star({[y]} @@ {[x]}) @@ {[x]}" then obtain p q r where pqr: "v=p@q@r" and "p\({[y]} @@ {[x]})" and q: "q \ star ({[y]} @@ {[x]})" and "r \ {[x]}" by (metis concE) then have "p = [y,x]" "r=[x]" by auto with pqr have a: "length v = 3+length q" by auto from q have b: "length q mod 2 = 0" apply(induct q rule: star_induct) by (auto) from a b show "length v mod 2 = Suc 0" by auto qed with assms(1,3) have fall: "(\hs. h = [x, x] @ hs) \ index h y = length h" by(auto) from assms(2) have "v \ lang (seq[Times (Atom y) (Atom x), Star(Times (Atom y) (Atom x))]) @@ lang (seq[Atom x])" by (auto simp: conc_def) then obtain a b where aa: "a \ lang (seq[Times (Atom y) (Atom x), Star(Times (Atom y) (Atom x))])" and "b \ lang (seq[Atom x])" and vab: "v = a @ b" by(erule concE) then have bb: "b=[x]" by auto from aa have lena: "length a > 0" by auto from TS_yxyx'[OF assms(1) aa fall] have stars: "T_on' (rTS h0) ([x, y], h) (a @ b) = length a - 1 + T_on' (rTS h0) ([x, y],rev a @ h) b" and history: "(\hs. rev a @ h = [x, y] @ hs)" and state: "config' (rTS h0) ([x, y], h) a = ([x, y], rev a @ h)" by auto have suffix: "T_on' (rTS h0) ( [x, y],rev a @ h) b = 0" and suState: "config' (rTS h0) ([x, y], rev a @ h) b = ([x,y], rev b @ (rev a @ h))" unfolding bb using TS_x' by auto from stars suffix have "T_on' (rTS h0) ([x, y], h) (a @ b) = length a - 1" by auto then have whatineed: "T_on' (rTS h0) ([x, y], h) v = (length v - 2)" using vab bb by auto have conf: "config' (rTS h0) ([x, y], h) v = ([x, y], rev v @ h)" by(simp add: vab config'_append2 state suState) from history obtain hs' where "rev a @ h = [x, y] @ hs'" by auto then obtain hs2 where reva: "rev a @ h = x # hs2" by auto show ?thesis using whatineed apply(auto) using conf apply(simp) by(simp add: reva vab bb) qed lemma TS_c1'': assumes "x \ y" "{x, y} = {x0, y0}" "TS_inv s x [x0, y0]" "set qs \ {x, y}" "qs \ lang (seq [Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom x])" shows "TS_inv (config'_rand (embed (rTS h0)) s qs) (last qs) [x0, y0] \ T_on_rand' (embed (rTS h0)) s qs = (length qs - 2)" proof - from assms(1,2) have kas: "(x0=x \ y0=y) \ (y0=x \ x0=y)" by(auto) then obtain h where S: "s = return_pmf ([x,y],h)" and h: "h = [] \ (\hs. h = [x, x] @ hs)" apply(rule disjE) using assms(1,3) unfolding TS_inv_def by(auto) have l: "qs \ []" using assms by auto { fix x y qs h0 fix h:: "nat list" assume A: "x \ y" and B: "qs \ lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom x])" and C: "h = [] \ (\hs. h = [x, x] @ hs)" then have C': "(\hs. h = [x, x] @ hs) \ h = [x] \ h = []" by blast from B have lqs: "last qs = x" using assms(5) by(auto simp add: conc_def) have "TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) qs) (last qs) [x, y] \ T_on_rand' (embed (rTS h0)) (return_pmf ([x, y], h)) qs = length qs - 2" apply(simp only: T_on'_embed[symmetric] config'_embed) using ts_c'[OF A B C'] A lqs unfolding TS_inv'_det by auto } note c1=this show ?thesis unfolding S using kas apply(rule disjE) apply(simp only:) apply(rule c1) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) apply(simp only:) apply(subst TS_inv_sym[of y x x y]) using assms(1) apply(simp) apply(blast) defer apply(rule c1) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) using last_in_set l assms(4) by blast qed lemma ts_c2': assumes "x \ y" "qs \ lang (seq[Atom x, Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom x])" "(\hs. h = [x, x] @ hs) \ h = []" shows "T_on' (rTS h0) ([x, y], h) qs = (length qs - 3) \ config' (rTS h0) ([x,y], h) qs = ([x,y],rev qs@h) \ (\hs. (rev qs @ h) = [x,x]@hs)" proof - from assms(2) obtain v where qs: "qs = [x]@v" and V: "v\lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom x])" by(auto simp add: conc_assoc) from assms(3) have 3: "(\hs. x#h = [x, x] @ hs) \ x#h = [x] \ x#h = []" by auto from ts_c'[OF assms(1) V 3] have T: "T_on' (rTS h0) ([x, y], x#h) v = length v - 2" and C: "config' (rTS h0) ([x, y], x#h) v = ([x, y], rev v @ x#h)" and H: "(\hs. rev v @ x#h = [x, x] @ hs)" by auto have t: "t\<^sub>p [x, y] x (fst (snd (rTS h0) ([x, y], h) x)) = 0" by (simp add: step_def rTS_def TS_step_d_def t\<^sub>p_def) have c: "Partial_Cost_Model.Step (rTS h0) ([x, y], h) x = ([x,y], x#h)" by (simp add: Step_def rTS_def TS_step_d_def step_def) show ?thesis unfolding qs apply(safe) apply(simp add: T_on'_append T c t) apply(simp add: config'_rand_append C c) using H by simp qed lemma TS_c2'': assumes "x \ y" "{x, y} = {x0, y0}" "TS_inv s x [x0, y0]" "set qs \ {x, y}" "qs \ lang (seq [Atom x, Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom x])" shows "TS_inv (config'_rand (embed (rTS h0)) s qs) (last qs) [x0, y0] \ T_on_rand' (embed (rTS h0)) s qs = (length qs - 3)" proof - from assms(1,2) have kas: "(x0=x \ y0=y) \ (y0=x \ x0=y)" by(auto) then obtain h where S: "s = return_pmf ([x,y],h)" and h: "h = [] \ (\hs. h = [x, x] @ hs)" apply(rule disjE) using assms(1,3) unfolding TS_inv_def by(auto) have l: "qs \ []" using assms by auto { fix x y qs h0 fix h:: "nat list" assume A: "x \ y" and B: "qs \ lang (seq[Atom x, Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom x])" and C: "h = [] \ (\hs. h = [x, x] @ hs)" from B have lqs: "last qs = x" using assms(5) by(auto simp add: conc_def) from C have C': "(\hs. h = [x, x] @ hs) \ h = []" by blast have "TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) qs) (last qs) [x, y] \ T_on_rand' (embed (rTS h0)) (return_pmf ([x, y], h)) qs = length qs - 3" apply(simp only: T_on'_embed[symmetric] config'_embed) using ts_c2'[OF A B C'] A lqs unfolding TS_inv'_det by auto } note c2=this show ?thesis unfolding S using kas apply(rule disjE) apply(simp only:) apply(rule c2) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) apply(simp only:) apply(subst TS_inv_sym[of y x x y]) using assms(1) apply(simp) apply(blast) defer apply(rule c2) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) using last_in_set l assms(4) by blast qed lemma TS_c': assumes "x \ y" "h = [] \ (\hs. h = [x, x] @ hs)" "qs \ lang (seq [Plus (Atom x) rexp.One, Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom x])" shows "T_on' (rTS h0) ([x, y], h) qs \ 2 * T\<^sub>p [x, y] qs (OPT2 qs [x, y]) \ TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x,y]" proof - obtain u v where uu: "u \ lang (Plus (Atom x) One)" and vv: "v \ lang (seq[Times (Atom y) (Atom x), Star (Times (Atom y) (Atom x)), Atom x])" and qsuv: "qs = u @ v" using assms(3) by (auto simp: conc_def) from TS_xr'[OF assms(1) uu assms(2)] have T_pre: "T_on' (rTS h0) ([x, y], h) (u@v) = T_on' (rTS h0) ([x, y], rev u @ h) v" and fall': "(\hs. (rev u @ h) = [x, x] @ hs) \ (rev u @ h) = [x] \ (rev u @ h)=[]" and conf': "config' (rTS h0) ([x, y], h) (u @ v) = config' (rTS h0) ([x, y], rev u @ h) v" by auto with assms uu have fall: "(\hs. (rev u @ h) = [x, x] @ hs) \ index (rev u @ h) y = length (rev u @ h)" by(auto) from ts_c'[OF assms(1) vv fall'] have T_star: "T_on' (rTS h0) ([x, y], rev u @ h) v = (length v - 2)" and inv1: "config' (rTS h0) ([x, y], (rev u @ h)) v = ([x, y], rev v @ rev u @ h)" and inv2: "(\hs. rev v @ rev u @ h = [x, x] @ hs)" by auto from T_pre T_star qsuv have TS: "T_on' (rTS h0) ([x, y], h) qs = (length v - 2)" by metis (* OPT *) from uu have uuu: "u=[] \ u=[x]" by auto from vv have vvv: "v \ lang (seq [Atom y, Atom x, Star (Times (Atom y) (Atom x)), Atom x])" by(auto simp: conc_def) have OPT: "T\<^sub>p [x,y] qs (OPT2 qs [x,y]) = (length v) div 2" apply(rule OPT2_C) by(fact)+ have lqs: "last qs = x" using assms(3) by force have conf: "config' (rTS h0) ([x, y], h) qs = ([x, y], rev qs @ h)" by(simp add: qsuv conf' inv1) then have conf: "TS_inv' (config' (rTS h0) ([x, y], h) qs) (last qs) [x,y]" apply(simp add: lqs) apply( subst TS_inv'_det) using inv2 qsuv by(simp) show ?thesis unfolding TS OPT by (auto simp add: conf) qed subsubsection "xx" lemma request_first: "x\y \ Step (rTS h) ([x, y], is) x = ([x,y],x#is)" unfolding rTS_def Step_def by(simp add: split_def TS_step_d_def step_def) lemma ts_d': "qs \ Lxx x y \ x \ y \ h = [] \ (\hs. h = [x, x] @ hs) \ qs \ lang (seq [Atom x, Atom x]) \ T_on' (rTS h0) ([x, y], h) qs = 0 \ TS_inv' (config' (rTS h0) ([x, y], h) qs) x [x,y]" proof - assume xny: "x \ y" assume "qs \ lang (seq [Atom x, Atom x])" then have xx: "qs = [x,x]" by auto from xny have TS: "T_on' (rTS h0) ([x, y], h) qs = 0" unfolding xx by(auto simp add: Step_def step_def oneTS_steps rTS_def t\<^sub>p_def) from xny have "config' (rTS h0) ([x, y], h) qs = ([x, y], x # x # h) " by(auto simp add: xx Step_def rTS_def oneTS_steps step_def) then have " TS_inv' (config' (rTS h0) ([x, y], h) qs) x [x, y]" by(simp add: TS_inv'_det) with TS show ?thesis by simp qed lemma TS_d': assumes xny: "x \ y" and "h = [] \ (\hs. h = [x, x] @ hs)" and qsis: "qs \ lang (seq [Atom x, Atom x])" shows "T_on' (rTS h0) ([x,y],h) qs \ 2 * T\<^sub>p [x, y] qs (OPT2 qs [x, y]) " and "TS_inv' (config' (rTS h0) ([x,y],h) qs) (last qs) [x, y]" and "T_on' (rTS h0) ([x,y],h) qs = 0" proof - from qsis have xx: "qs = [x,x]" by auto show TS: "T_on' (rTS h0) ([x,y],h) qs = 0" using assms(1) by (auto simp add: xx t\<^sub>p_def rTS_def Step_def oneTS_steps step_def) then show "T_on' (rTS h0) ([x,y],h) qs \ 2 * T\<^sub>p [x, y] qs (OPT2 qs [x, y])" by simp show "TS_inv' (config' (rTS h0) ([x,y],h) qs) (last qs) [x, y]" unfolding TS_inv_def by(simp add: xx request_first[OF xny]) qed lemma TS_d'': assumes "x \ y" "{x, y} = {x0, y0}" "TS_inv s x [x0, y0]" "set qs \ {x, y}" "qs \ lang (seq [Atom x, Atom x])" shows "TS_inv (config'_rand (embed (rTS h0)) s qs) (last qs) [x0, y0] \ T_on_rand' (embed (rTS h0)) s qs = 0" proof - from assms(1,2) have kas: "(x0=x \ y0=y) \ (y0=x \ x0=y)" by(auto) then obtain h where S: "s = return_pmf ([x,y],h)" and h: "h = [] \ (\hs. h = [x, x] @ hs)" apply(rule disjE) using assms(1,3) unfolding TS_inv_def by(auto) have l: "qs \ []" using assms by auto { fix x y qs h0 fix h:: "nat list" assume A: "x \ y" and B: "qs \ lang (seq [Atom x, Atom x])" and C: "h = [] \ (\hs. h = [x, x] @ hs)" from B have lqs: "last qs = x" using assms(5) by(auto simp add: conc_def) have "TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) qs) (last qs) [x, y] \ T_on_rand' (embed (rTS h0)) (return_pmf ([x, y], h)) qs = 0" apply(simp only: T_on'_embed[symmetric] config'_embed) using TS_d'[OF A C B ] A lqs unfolding TS_inv'_det by auto } note d=this show ?thesis unfolding S using kas apply(rule disjE) apply(simp only:) apply(rule d) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) apply(simp only:) apply(subst TS_inv_sym[of y x x y]) using assms(1) apply(simp) apply(blast) defer apply(rule d) using assms apply(simp) using assms apply(simp add: conc_assoc) using h apply(simp) using last_in_set l assms(4) by blast qed subsection "Phase Partitioning" lemma D': assumes "\' \ Lxx x y" and "x \ y" and "TS_inv' ([x, y], h) x [x, y]" shows "T_on' (rTS h0) ([x, y], h) \' \ 2 * T\<^sub>p [x, y] \' (OPT2 \' [x, y]) \ TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) \') (last \') [x, y]" proof - from config'_embed have " config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) \' = return_pmf (Partial_Cost_Model.config' (rTS h0) ([x, y], h) \')" by blast then have L: "TS_inv (config'_rand (embed (rTS h0)) (return_pmf ([x, y], h)) \') (last \') [x, y] = TS_inv' (config' (rTS h0) ([x, y], h) \') (last \') [x, y]" by auto from assms(3) have h: "h = [] \ (\hs. h = [x, x] @ hs)" by(auto simp add: TS_inv'_det) have "T_on' (rTS h0) ([x, y], h) \' \ 2 * T\<^sub>p [x, y] \' (OPT2 \' [x, y]) \ TS_inv' (config' (rTS h0) ([x, y], h) \') (last \') [x, y]" apply(rule LxxE[OF assms(1)]) using TS_d'[OF assms(2) h, of "\'"] apply(simp) using TS_b'[OF assms(2) h] apply(simp) using TS_c'[OF assms(2) h] apply(simp) using TS_a'[OF assms(2) h] apply fast done then show ?thesis using L by auto qed theorem TS_OPT2': "(x::nat) \ y \ set \ \ {x,y} \ T\<^sub>p_on (rTS []) [x,y] \ \ 2 * real (T\<^sub>p_opt [x,y] \) + 2" apply(subst T_on_embed) apply(rule Phase_partitioning_general[where P=TS_inv]) apply(simp) apply(simp) apply(simp) apply(simp add: TS_inv_def rTS_def) proof (goal_cases) case (1 a b \' s) from 1(6) obtain h hist' where s: "s = return_pmf ([a, b], h)" and "h = [] \ h = [a,a]@hist'" unfolding TS_inv_def apply(cases "a=hd [x,y]") apply(simp) using 1 apply fast apply(simp) using 1 by blast from 1 have xyab: "TS_inv' ([a, b], h) a [x, y] = TS_inv' ([a, b], h) a [a, b]" by(auto simp add: TS_inv'_det) with 1(6) s have inv: "TS_inv' ([a, b], h) a [a, b]" by simp from \\' \ Lxx a b\ have "\' \ []" using Lxx1 by fastforce then have l: "last \' \ {x,y}" using 1(5,7) last_in_set by blast show ?case unfolding s T_on'_embed[symmetric] using D'[OF 1(3,4) inv, of "[]"] apply(safe) apply linarith using TS_inv_sym[OF 1(4,5)] l apply blast done qed subsection "TS is pairwise" lemma config'_distinct[simp]: shows "distinct (fst (config' A S qs)) = distinct (fst S)" apply (induct qs rule: rev_induct) by(simp_all add: config'_snoc Step_def split_def distinct_step) lemma config'_set[simp]: shows "set (fst (config' A S qs)) = set (fst S)" apply (induct qs rule: rev_induct) by(simp_all add: config'_snoc Step_def split_def set_step) lemma s_TS_append: "i\length as \s_TS init h (as@bs) i = s_TS init h as i" by (simp add: s_TS_def) lemma s_TS_distinct: "distinct init \ i distinct (fst (TSdet init h qs i))" by(simp_all add: config_config_distinct) lemma othersdontinterfere: "distinct init \ i < length qs \ a\set init \ b\set init \ set qs \ set init \ qs!i\{a,b} \ a < b in s_TS init h qs i \ a < b in s_TS init h qs (Suc i)" apply(simp add: s_TS_def split_def take_Suc_conv_app_nth config_append Step_def step_def) apply(subst x_stays_before_y_if_y_not_moved_to_front) apply(simp_all add: config_config_distinct config_config_set) by(auto simp: rTS_def TS_step_d_def) lemma TS_mono: fixes l::nat assumes 1: "x < y in s_TS init h xs (length xs)" and l_in_cs: "l \ length cs" and firstocc: "\j y" and "x \ set cs" and di: "distinct init" and inin: "set (xs @ cs) \ set init" shows "x < y in s_TS init h (xs@cs) (length (xs)+l)" proof - from before_in_setD2[OF 1] have y: "y : set init" unfolding s_TS_def by(simp add: config_config_set) from before_in_setD1[OF 1] have x: "x : set init" unfolding s_TS_def by(simp add: config_config_set) { fix n assume "n\l" then have "x < y in s_TS init h ((xs)@cs) (length (xs)+n)" proof(induct n) case 0 show ?case apply (simp only: s_TS_append ) using 1 by(simp) next case (Suc n) then have n_lt_l: "n \before the request to y, x is in front of y\ then show "x < y in s_TS init h (xs@cs) (length (xs)+l)" by blast qed lemma step_no_action: "step s q (0,[]) = s" unfolding step_def mtf2_def by simp lemma s_TS_set: "i \ length qs \ set (s_TS init h qs i) = set init" apply(induct i) apply(simp add: s_TS_def ) apply(simp add: s_TS_def TSdet_Suc) by(simp add: split_def rTS_def Step_def step_def) lemma count_notin2: "count_list xs x = 0 \ x \ set xs" by (simp add: count_list_0_iff) lemma mtf2_q_passes: assumes "q \ set xs" "distinct xs" and "index xs q - n \ index xs x" "index xs x < index xs q" shows "q < x in (mtf2 n q xs)" proof - from assms have "index xs q < length xs" by auto with assms(4) have ind_x: "index xs x < length xs" by auto then have xinxs: "x\set xs" using index_less_size_conv by metis have B: "index (mtf2 n q xs) q = index xs q - n" apply(rule mtf2_q_after) by(fact)+ also from ind_x mtf2_forward_effect3'[OF assms] have A: "\ < index (mtf2 n q xs) x" by auto finally show ?thesis unfolding before_in_def using xinxs by force qed lemma twotox: assumes "count_list bs y \ 1" and "distinct init" and "x \ set init" and "y : set init" and "x \ set bs" and "x\y" shows "x < y in s_TS init h (as@[x]@bs@[x]) (length (as@[x]@bs@[x]))" proof - have aa: "snd (TSdet init h ((as @ x # bs) @ [x]) (Suc (length as + length bs))) = rev (take (Suc (length as + length bs)) ((as @ x # bs) @ [x])) @ h" apply(rule sndTSdet) by(simp) then have aa': "snd (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs))) = rev (take (Suc (length as + length bs)) ((as @ x # bs) @ [x])) @ h" by auto have lasocc_x: "index (snd (TSdet init h ((as @ x # bs) @ [x]) (Suc (length as + length bs)))) x = length bs" unfolding aa apply(simp add: del: config'.simps) using assms(5) by(simp add: index_append) then have lasocc_x': "(index (snd (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) x) = length bs" by auto let ?sincelast = "take (length bs) (snd (TSdet init h ((as @ x # bs) @ [x]) (Suc (length as + length bs))))" have sl: "?sincelast = rev bs" unfolding aa by(simp) let ?S = "{xa. xa < x in fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs))) \ count_list ?sincelast xa \ 1}" have y: "y \ ?S \ ~ y < x in s_TS init h (as @ x # bs @ [x]) (Suc (length as + length bs))" unfolding sl unfolding s_TS_def using assms(1) by(simp del: config'.simps) have eklr: "length (as@[x]@bs@[x]) = Suc (length (as@[x]@bs))" by simp have 1: "s_TS init h (as@[x]@bs@[x]) (length (as@[x]@bs@[x])) = fst (Partial_Cost_Model.Step (rTS h) (TSdet init h (as @ [x] @ bs @ [x]) (length (as @ [x] @ bs))) ((as @ [x] @ bs @ [x]) ! length (as @ [x] @ bs)))" unfolding s_TS_def unfolding eklr apply(subst TSdet_Suc) by(simp_all add: split_def) have brrr: "x\ set (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs))))" apply(subst s_TS_set[unfolded s_TS_def]) apply(simp) by fact have ydrin: "y\set (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs))))" apply(subst s_TS_set[unfolded s_TS_def]) apply(simp) by fact have dbrrr: "distinct (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs))))" apply(subst s_TS_distinct[unfolded s_TS_def]) using assms(2) by(simp_all) show ?thesis proof (cases "y < x in s_TS init h (as @ x # bs @ [x]) (Suc (length as + length bs))") case True with y have yS: "y\?S" by auto then have minsteps: "Min (index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) ` ?S) \ index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) y" by auto let ?entf = "index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) x - Min (index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) ` ?S)" from minsteps have br: " index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) x - (?entf) \ index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) y" by presburger have brr: "index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) y < index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) x" using True unfolding before_in_def s_TS_def by auto from br brr have klo: " index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) x - (?entf) \ index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) y \ index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) y < index (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))) x" by metis let ?result ="(mtf2 ?entf x (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))))" have whatsthat: "s_TS init h (as @ [x] @ bs @ [x]) (length (as @ [x] @ bs @ [x])) = ?result" unfolding 1 apply(simp add: split_def step_def rTS_def Step_def TS_step_d_def del: config'.simps) apply(simp add: nth_append del: config'.simps) using lasocc_x'[unfolded rTS_def] aa'[unfolded rTS_def] apply(simp add: del: config'.simps) using yS[unfolded sl rTS_def] by auto have ydrinee: " y \ set (mtf2 ?entf x (fst (TSdet init h (as @ x # bs @ [x]) (Suc (length as + length bs)))))" apply(subst set_mtf2) apply(subst s_TS_set[unfolded s_TS_def]) apply(simp) by fact show ?thesis unfolding whatsthat apply(rule mtf2_q_passes) by(fact)+ next case False then have 2: "x < y in s_TS init h (as @ x # bs @ [x]) (Suc (length as + length bs))" using brrr ydrin not_before_in assms(6) unfolding s_TS_def by metis { fix e have "x < y in mtf2 e x (s_TS init h (as @ x # bs @ [x]) (Suc (length as + length bs)))" apply(rule x_stays_before_y_if_y_not_moved_to_front) unfolding s_TS_def apply(fact)+ using assms(6) apply(simp) using 2 unfolding s_TS_def by simp } note bratz=this show ?thesis unfolding 1 apply(simp add: TSnopaid split_def Step_def s_TS_def TS_step_d_def step_def nth_append del: config'.simps) using bratz[unfolded s_TS_def] by simp qed qed lemma count_drop: "count_list (drop n cs) x \ count_list cs x" proof - have "count_list cs x = count_list (take n cs @ drop n cs) x" by auto also have "\ = count_list (take n cs) x + count_list (drop n cs) x" by (rule count_list_append) also have "\ \ count_list (drop n cs) x" by auto finally show ?thesis . qed lemma count_take_less: assumes "n\m" shows "count_list (take n cs) x \ count_list (take m cs) x" proof - from assms have "count_list (take n cs) x = count_list (take n (take m cs)) x" by auto also have "\ \ count_list (take n (take m cs) @ drop n (take m cs)) x" by (simp) also have "\ = count_list (take m cs) x" by(simp only: append_take_drop_id) finally show ?thesis . qed lemma count_take: "count_list (take n cs) x \ count_list cs x" proof - have "count_list cs x = count_list (take n cs @ drop n cs) x" by auto also have "\ = count_list (take n cs) x + count_list (drop n cs) x" by (rule count_list_append) also have "\ \ count_list (take n cs) x" by auto finally show ?thesis . qed lemma casexxy: assumes "\=as@[x]@bs@[x]@cs" and "x \ set cs" and "set cs \ set init" and "x \ set init" and "distinct init" and "x \ set bs" and "set as \ set init" and "set bs \ set init" shows "(%i. i (\jcs!i) \ cs!i\x \ (cs!i) \ set bs \ x < (cs!i) in (s_TS init h \ (length (as@[x]@bs@[x]) + i+1))) i" proof (rule infinite_descent[where P="(%i. i (\jcs!i) \ cs!i\x \ (cs!i) \ set bs \ x < (cs!i) in (s_TS init h \ (length (as@[x]@bs@[x]) + i+1)))"], goal_cases) case (1 i) let ?y = "cs!i" from 1 have i_in_cs: "i < length cs" and firstocc: "(\j cs ! i)" and ynx: "cs ! i \ x" and ynotinbs: "cs ! i \ set bs" and y_before_x': "~x < cs ! i in s_TS init h \ (length (as @ [x] @ bs @ [x]) + i+1)" by auto have ss: "set (s_TS init h \ (length (as @ [x] @ bs @ [x]) + i+1)) = set init" using assms(1) i_in_cs by(simp add: s_TS_set) then have "cs ! i \ set (s_TS init h \ (length (as @ [x] @ bs @ [x]) + i+1))" unfolding ss using assms(3) i_in_cs by fastforce moreover have "x : set (s_TS init h \ (length (as @ [x] @ bs @ [x]) + i+1))" unfolding ss using assms(4) by fastforce \ \after the request to y, y is in front of x\ ultimately have y_before_x_Suct3: "?y < x in s_TS init h \ (length (as @ [x] @ bs @ [x]) + i+1)" using y_before_x' ynx not_before_in by metis from ynotinbs have yatmostonceinbs: "count_list bs (cs!i) \ 1" by simp let ?y = "cs!i" have yininit: "?y \ set init" using assms(3) i_in_cs by fastforce { fix y assume "y \ set init" assume "x\y" assume "count_list bs y \ 1" then have "x < y in s_TS init h (as@[x]@bs@[x]) (length (as@[x]@bs@[x]))" apply(rule twotox) by(fact)+ } note xgoestofront=this with yatmostonceinbs ynx yininit have zeitpunktt2: "x < ?y in s_TS init h (as@[x]@bs@[x]) (length (as@[x]@bs@[x]))" by blast have "i \ length cs" using i_in_cs by auto have x_before_y_t3: "x < ?y in s_TS init h ((as@[x]@bs@[x])@cs) (length (as@[x]@bs@[x])+i)" apply(rule TS_mono) apply(fact)+ using assms by simp \ \so x and y swap positions when y is requested, that means that y was inserted infront of some elment z (which cannot be x, has only been requested at most once since last request of y but is in front of x)\ \ \first show that y must have been requested in as\ have "snd (TSdet init h (as @ [x] @ bs @ [x] @ cs) (length (as @ [x] @ bs @ [x]) + i)) = rev (take (length (as @ [x] @ bs @ [x]) + i) (as @ [x] @ bs @ [x] @ cs)) @ h" apply(rule sndTSdet) using i_in_cs by simp also have "\ = (rev (take i cs)) @ [x] @ (rev bs) @ [x] @ (rev as) @ h" by simp finally have fstTS_t3: "snd (TSdet init h (as @ [x] @ bs @ [x] @ cs) (length (as @ [x] @ bs @ [x]) + i)) = (rev (take i cs)) @ [x] @ (rev bs) @ [x] @ (rev as) @ h" . then have fstTS_t3': "(snd (TSdet init h \ (Suc (Suc (length as + length bs + i))))) = (rev (take i cs)) @ [x] @ (rev bs) @ [x] @ (rev as) @ h" using assms(1) by auto let ?is = "snd (TSdet init h (as @ [x] @ bs @ [x] @ cs) (length (as @ [x] @ bs @ [x]) + i))" let ?is' = "snd (config (rTS h) init (as @ [x] @ bs @ [x] @ (take i cs)))" let ?s = "fst (TSdet init h (as @ [x] @ bs @ [x] @ cs) (length (as @ [x] @ bs @ [x]) + i))" let ?s' = "fst (config (rTS h) init (as @ [x] @ bs @ [x] @ (take i cs)))" let ?s_Suct3="s_TS init h (as @ [x] @ bs @ [x] @ cs) (length (as @ [x] @ bs @ [x]) + i+1)" let ?S = "{xa. (xa < (as @ [x] @ bs @ [x] @ cs) ! (length (as @ [x] @ bs @ [x]) + i) in ?s \ count_list (take (index ?is ((as @ [x] @ bs @ [x] @ cs) ! (length (as @ [x] @ bs @ [x]) + i))) ?is) xa \ 1) }" let ?S' = "{xa. (xa < (as @ [x] @ bs @ [x] @ cs) ! (length (as @ [x] @ bs @ [x]) + i) in ?s' \ count_list (take (index ?is' ((cs!i))) ?is') xa \ 1) }" have isis': "?is = ?is'" by(simp) have ss': "?s = ?s'" by(simp) then have SS': "?S = ?S'" using i_in_cs by(simp add: nth_append) (* unfold TSdet once *) have once: "TSdet init h (as @ x # bs @ x # cs) (Suc (Suc (Suc (length as + length bs + i)))) = Step (rTS h) (config\<^sub>p (rTS h) init (as @ x # bs @ x # take i cs)) (cs ! i)" apply(subst TSdet_Suc) using i_in_cs apply(simp) by(simp add: nth_append) have aha: "(index ?is (cs ! i) \ length ?is) \ ?S \ {}" proof (rule ccontr, goal_cases) case 1 then have "(index ?is (cs ! i) = length ?is) \ ?S = {}" by(simp) then have alters: "(index ?is' (cs ! i) = length ?is') \ ?S' = {}" apply(simp only: SS') by(simp only: isis') \ \wenn (cs ! i) noch nie requested wurde, dann kann es gar nicht nach vorne gebracht werden! also widerspruch mit @{text y_before_x'}\ have "?s_Suct3 = fst (config (rTS h) init ((as @ [x] @ bs @ [x]) @ (take (i+1) cs)))" unfolding s_TS_def apply(simp only: length_append) apply(subst take_append) apply(subst take_append) apply(subst take_append) apply(subst take_append) by(simp) also have "\ = fst (config (rTS h) init (((as @ [x] @ bs @ [x]) @ (take i cs)) @ [cs!i]))" using i_in_cs by(simp add: take_Suc_conv_app_nth) also have "\ = step ?s' ?y (0, [])" proof (cases "index ?is' (cs ! i) = length ?is'") case True show ?thesis apply(subst config_append) using i_in_cs apply(simp add: rTS_def Step_def split_def nth_append) apply(subst TS_step_d_def) apply(simp only: True[unfolded rTS_def,simplified]) by(simp) next case False with alters have S': "?S' = {}" by simp have 1 : "{xa. xa < cs ! i in fst (Partial_Cost_Model.config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs)) \ count_list (take (index (snd (Partial_Cost_Model.config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs))) (cs ! i)) (snd (Partial_Cost_Model.config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs)))) xa \ 1} = {}" using S' by(simp add: rTS_def nth_append) show ?thesis apply(subst config_append) using i_in_cs apply(simp add: rTS_def Step_def split_def nth_append) apply(subst TS_step_d_def) apply(simp only: 1 Let_def) by(simp) qed finally have "?s_Suct3 = step ?s ?y (0, [])" using ss' by simp then have e: "?s_Suct3 = ?s" by(simp only: step_no_action) from x_before_y_t3 have "x < cs ! i in ?s_Suct3" unfolding e unfolding s_TS_def by simp with y_before_x' show "False" unfolding assms(1) by auto qed then have aha': "index (snd (TSdet init h (as @ x # bs @ x # cs) (Suc (Suc (length as + length bs + i))))) (cs ! i) \ length (snd (TSdet init h (as @ x # bs @ x # cs) (Suc (Suc (length as + length bs + i)))))" and aha2: "?S \ {}" by auto from fstTS_t3' assms(1) have is_: "?is = (rev (take i cs)) @ [x] @ (rev bs) @ [x] @ (rev as) @ h" by auto have minlencsi: " min (length cs) i = i" using i_in_cs by linarith let ?lastoccy="(index (rev (take i cs) @ x # rev bs @ x # rev as @ h) (cs ! i))" have "?y \ set (rev (take i cs))" using firstocc by (simp add: in_set_conv_nth) then have lastoccy: "?lastoccy \ i + 1 + length bs + 1" using ynx ynotinbs minlencsi by(simp add: index_append) (* x is not in S, because it is requested at least twice since the last request to y*) have x_nin_S: "x\?S" using is_ apply(simp add: split_def nth_append del: config'.simps) proof (goal_cases) case 1 have " count_list (take ?lastoccy (rev (take i cs))) x \ count_list (drop (length cs - i) (rev cs)) x" by (simp add: count_take rev_take) also have "\ \ count_list (rev cs) x" by (meson count_drop) also have "\ = 0" using assms(2) by(simp) finally have " count_list (take ?lastoccy (rev (take i cs))) x = 0" by auto have" 2 \ count_list ([x] @ rev bs @ [x]) x " by(simp) also have "\ = count_list (take (1 + length bs + 1) (x # rev bs @ x # rev as @ h)) x" by auto also have "\ \ count_list (take (?lastoccy - i) (x # rev bs @ x # rev as @ h)) x" apply(rule count_take_less) using lastoccy by linarith also have "\ \ count_list (take ?lastoccy (rev (take i cs))) x + count_list (take (?lastoccy - i) (x # rev bs @ x # rev as @ h)) x" by auto finally show ?case by(simp add: minlencsi) qed have "Min (index ?s ` ?S) \ (index ?s ` ?S)" apply(rule Min_in) using aha2 by (simp_all) then obtain z where zminimal: "index ?s z = Min (index ?s ` ?S)"and z_in_S: "z \ ?S" by auto then have bef: "z < (as @ [x] @ bs @ [x] @ cs) ! (length (as @ [x] @ bs @ [x]) + i) in ?s" and "count_list (take (index ?is ((as @ [x] @ bs @ [x] @ cs) ! (length (as @ [x] @ bs @ [x]) + i))) ?is) z \ 1" by(blast)+ with zminimal have zbeforey: "z < cs ! i in ?s" and zatmostonce: "count_list (take (index ?is (cs ! i)) ?is) z \ 1" and isminimal: "index ?s z = Min (index ?s ` ?S)" by(simp_all add: nth_append) have elemins: "z \ set ?s" unfolding before_in_def by (meson zbeforey before_in_setD1) then have zininit: "z \ set init" using i_in_cs by(simp add: s_TS_set[unfolded s_TS_def] del: config'.simps) from zbeforey have zbeforey_ind: "index ?s z < index ?s ?y" unfolding before_in_def by auto then have el_n_y: "z \ ?y" by auto have el_n_x: "z \ x" using x_nin_S z_in_S by blast (* and because it is JUST before that element, z must be before x *) { fix s q have TS_step_d2: "TS_step_d s q = (let V\<^sub>r={x. x < q in fst s \ count_list (take (index (snd s) q) (snd s)) x \ 1} in ((if index (snd s) q \ length (snd s) \ V\<^sub>r \ {} then index (fst s) q - Min ( (index (fst s)) ` V\<^sub>r) else 0,[]),q#(snd s)))" unfolding TS_step_d_def apply(cases "index (snd s) q < length (snd s)") using index_le_size apply(simp split: prod.split) apply blast by(auto simp add: index_less_size_conv split: prod.split) } note alt_chara=this have iF: "(index (snd (config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs))) (cs ! i) \ length (snd (config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs))) \ {xa. xa < cs ! i in fst (config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs)) \ count_list (take (index (snd (config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs))) (cs ! i)) (snd (Partial_Cost_Model.config' (\s. h, TS_step_d) (init, h) (as @ x # bs @ x # take i cs)))) xa \ 1} \ {}) = True" using aha[unfolded rTS_def] ss' SS' by(simp add: nth_append) have "?s_Suct3 = fst (TSdet init h (as @ x # bs @ x # cs) (Suc (Suc (Suc (length as + length bs + i)))))" by(auto simp add: s_TS_def) also have "\ = step ?s ?y (index ?s ?y - Min (index ?s ` ?S), [])" apply(simp only: once[unfolded assms(1)]) apply(simp add: Step_def split_def rTS_def del: config'.simps) apply(subst alt_chara) apply(simp only: Let_def ) apply(simp only: iF) by(simp add: nth_append) finally have "?s_Suct3 = step ?s ?y (index ?s ?y - Min (index ?s ` ?S), [])" . with isminimal have state_dannach: "?s_Suct3 = step ?s ?y (index ?s ?y - index ?s z, [])" by presburger \ \so y is moved in front of z, that means:\ have yinfrontofz: "?y < z in s_TS init h \ (length (as @ [x] @ bs @ [x]) + i+1)" unfolding assms(1) state_dannach apply(simp add: step_def del: config'.simps) apply(rule mtf2_q_passes) using i_in_cs assms(5) apply(simp_all add: s_TS_distinct[unfolded s_TS_def] s_TS_set[unfolded s_TS_def]) using yininit apply(simp) using zbeforey_ind by simp have yins: "?y \ set ?s" using i_in_cs assms(3,5) apply(simp_all add: s_TS_set[unfolded s_TS_def] del: config'.simps) by fastforce have "index ?s_Suct3 ?y = index ?s z" and "index ?s_Suct3 z = Suc (index ?s z)" proof - let ?xs = "(fst (TSdet init h (as @ x # bs @ x # cs) (Suc (Suc (length as + length bs + i)))))" have setxs: "set ?xs = set init" apply(rule s_TS_set[unfolded s_TS_def]) using i_in_cs by auto then have yinxs: "cs ! i \ set ?xs" apply(simp add: setxs del: config'.simps) using assms(3) i_in_cs by fastforce have distinctxs: "distinct ?xs" apply(rule s_TS_distinct[unfolded s_TS_def]) using i_in_cs assms(5) by auto let ?n = "(index (fst (TSdet init h (as @ x # bs @ x # cs) (Suc (Suc (length as + length bs + i))))) (cs ! i) - index (fst (TSdet init h (as @ x # bs @ x # cs) (Suc (Suc (length as + length bs + i))))) z)" have "index (mtf2 ?n ?y ?xs) (?xs ! index ?xs ?y) = index ?xs ?y - ?n\ index ?xs ?y - ?n = index (mtf2 ?n ?y ?xs) (?xs ! index ?xs ?y )" apply(rule mtf2_forward_effect2) apply(fact) apply(fact) by simp then have "index (mtf2 ?n ?y ?xs) (?xs ! index ?xs ?y) = index ?xs ?y - ?n" by metis also have "\ = index ?s z" using zbeforey_ind by force finally have A: "index (mtf2 ?n ?y ?xs) (?xs ! index ?xs ?y) = index ?s z" . have aa: "index ?xs ?y - ?n \ index ?xs z" "index ?xs z < index ?xs ?y" apply(simp) using zbeforey_ind by fastforce from mtf2_forward_effect3'[OF yinxs distinctxs aa] have B: "index (mtf2 ?n ?y ?xs) z = Suc (index ?xs z)" using elemins yins by(simp add: nth_append split_def del: config'.simps) show "index ?s_Suct3 ?y = index ?s z" unfolding state_dannach apply(simp add: step_def nth_append del: config'.simps) using A yins by(simp add: nth_append del: config'.simps) show "index ?s_Suct3 z = Suc (index ?s z)" unfolding state_dannach apply(simp add: step_def nth_append del: config'.simps) using B yins by(simp add: nth_append del: config'.simps) qed then have are: "Suc (index ?s_Suct3 ?y) = index ?s_Suct3 z" by presburger from are before_in_def y_before_x_Suct3 el_n_x assms(1) have z_before_x: "z < x in ?s_Suct3" by (metis Suc_lessI not_before_in yinfrontofz) have xSuct3: "x\set ?s_Suct3" using assms(4) i_in_cs by(simp add: s_TS_set) have elSuct3: "z\set ?s_Suct3" using zininit i_in_cs by(simp add: s_TS_set) have xt3: "x\set ?s " apply(subst config_config_set) by fact note elt3=elemins have z_s: "z < x in ?s" proof(rule ccontr, goal_cases) case 1 then have "x < z in ?s" using not_before_in[OF xt3 elt3] el_n_x unfolding s_TS_def by blast then have "x < z in ?s_Suct3" apply (simp only: state_dannach) apply (simp only: step_def) apply(simp add: nth_append del: config'.simps) apply(rule x_stays_before_y_if_y_not_moved_to_front) apply(subst config_config_set) using i_in_cs assms(3) apply fastforce apply(subst config_config_distinct) using assms(5) apply fastforce apply(subst config_config_set) using assms(4) apply fastforce apply(subst config_config_set) using zininit apply fastforce using el_n_y apply(simp) by(simp) then show "False" using z_before_x not_before_in[OF xSuct3 elSuct3] by blast qed have mind: "(index ?is (cs ! i)) \ i + 1 + length bs + 1 " using lastoccy using i_in_cs fstTS_t3'[unfolded assms(1)] by(simp add: split_def nth_append del: config'.simps) have "count_list (rev (take i cs) @ [x] @ rev bs @ [x]) z= count_list (take (i + 1 + length bs + 1) ?is) z" unfolding is_ using el_n_x by(simp add: minlencsi) also from mind have "\ \ count_list (take (index ?is (cs ! i)) ?is) z" by(rule count_take_less) also have "\ \ 1" using zatmostonce by metis finally have aaa: "count_list (rev (take i cs) @ [x] @ rev bs @ [x]) z \ 1" . with el_n_x have "count_list bs z + count_list (take i cs) z \ 1" by(simp) moreover have "count_list (take (Suc i) cs) z = count_list (take i cs) z" using i_in_cs el_n_y by(simp add: take_Suc_conv_app_nth) ultimately have aaaa: "count_list bs z + count_list (take (Suc i) cs) z \ 1" by simp have z_occurs_once_in_cs: "count_list (take (Suc i) cs) z = 1" proof (rule ccontr, goal_cases) case 1 with aaaa have atmost1: "count_list bs z \ 1" and "count_list (take (Suc i) cs) z = 0" by force+ have yeah: "z \ set (take (Suc i) cs)" apply(rule count_notin2) by fact \ \now we know that x is in front of z after 2nd request to x, and that z is not requested any more, that means it stays behind x, which leads to a contradiction with @{text z_before_x}\ have xin123: "x \ set (s_TS init h ((as @ [x] @ bs @ [x]) @ (take (i+1) cs)) (length (as @ [x] @ bs @ [x]) + (i+1)))" using i_in_cs assms(4) by(simp add: s_TS_set) have zin123: "z \ set (s_TS init h ((as @ [x] @ bs @ [x]) @ (take (i+1) cs)) (length (as @ [x] @ bs @ [x]) + (i+1)))" using i_in_cs elemins by(simp add: s_TS_set del: config'.simps) have "x < z in s_TS init h ((as @ [x] @ bs @ [x]) @ (take (i+1) cs)) (length (as @ [x] @ bs @ [x]) + (i + 1))" apply(rule TS_mono) apply(rule xgoestofront) apply(fact) using el_n_x apply(simp) apply(fact) using i_in_cs apply(simp) using yeah i_in_cs length_take nth_mem apply (metis Suc_eq_plus1 Suc_leI min_absorb2) using set_take_subset assms(2) apply fast using assms i_in_cs apply(simp_all ) using set_take_subset by fast then have ge: "\ z < x in s_TS init h ((as @ [x] @ bs @ [x]) @ (take (i+1) cs)) (length (as @ [x] @ bs @ [x]) + (i+1))" using not_before_in[OF zin123 xin123] el_n_x by blast have " s_TS init h ((as @ [x] @ bs @ [x]) @ cs) (length (as @ [x] @ bs @ [x]) + (i+1)) = s_TS init h ((as @ [x] @ bs @ [x] @ (take (i+1) cs)) @ (drop (i+1) cs)) (length (as @ [x] @ bs @ [x]) + (i+1))" by auto also have "\ = s_TS init h (as @ [x] @ bs @ [x] @ (take (i+1) cs)) (length (as @ [x] @ bs @ [x]) + (i+1))" apply(rule s_TS_append) using i_in_cs by(simp) finally have aaa: " s_TS init h ((as @ [x] @ bs @ [x]) @ cs) (length (as @ [x] @ bs @ [x]) + (i+1)) = s_TS init h (as @ [x] @ bs @ [x] @ (take (i+1) cs)) (length (as @ [x] @ bs @ [x]) + (i+1))" . from ge z_before_x show "False" unfolding assms(1) using aaa by auto qed from z_occurs_once_in_cs have kinSuci: "z \ set (take (Suc i) cs)" by (metis One_nat_def count_notin n_not_Suc_n) then have zincs: "z\set cs" using set_take_subset by fast from z_occurs_once_in_cs obtain k where k_def: "k=index (take (Suc i) cs) z" by blast then have "k=index cs z" using kinSuci by (simp add: index_take_if_set) then have zcsk: "z = cs!k" using zincs by simp have era: " cs ! index (take (Suc i) cs) z = z" using kinSuci in_set_takeD index_take_if_set by fastforce have ki: "k \before the request to z, x is in front of z, analog zu oben, vllt generell machen?\ \ \element z does not occur between t1 and position k\ have z_notinbs: "cs ! k \ set bs" proof - from z_occurs_once_in_cs aaaa have "count_list bs z = 0" by auto then show ?thesis using zcsk count_notin2 by metis qed have "count_list bs z \ 1" using aaaa by linarith with xgoestofront[OF zininit el_n_x[symmetric]] have xbeforez: "x < z in s_TS init h (as @ [x] @ bs @ [x]) (length (as @ [x] @ bs @ [x]))" by auto obtain cs1 cs2 where v: "cs1 @ cs2 = cs" and cs1: "cs1 = take (Suc k) cs" and cs2: "cs2 = drop (Suc k) cs" by auto have z_firstocc: "\j cs ! k" and z_lastocc: "\j cs ! k" proof (safe, goal_cases) case (1 j) with ki i_in_cs have 2: "j < length (take k cs)" by auto have un1: "(take (Suc i) cs)!k = cs!k" apply(rule nth_take) using ki by auto have un2: "(take k cs)!j = cs!j" apply(rule nth_take) using 1(1) ki by auto from i_in_cs ki have f1: "k < length (take (Suc i) cs)" by auto then have "(take (Suc i) cs) = (take k (take (Suc i) cs)) @ (take (Suc i) cs)!k # (drop (Suc k) (take (Suc i) cs))" by(rule id_take_nth_drop) also have "(take k (take (Suc i) cs)) = take k cs" using i_in_cs ki by (simp add: min_def) also have "... = (take j (take k cs)) @ (take k cs)!j # (drop (Suc j) (take k cs))" using 2 by(rule id_take_nth_drop) finally have "take (Suc i) cs = (take j (take k cs)) @ [(take k cs)!j] @ (drop (Suc j) (take k cs)) @ [(take (Suc i) cs)!k] @ (drop (Suc k) (take (Suc i) cs))" by(simp) then have A: "take (Suc i) cs = (take j (take k cs)) @ [cs!j] @ (drop (Suc j) (take k cs)) @ [cs!k] @ (drop (Suc k) (take (Suc i) cs))" unfolding un1 un2 by simp have "count_list ((take j (take k cs)) @ [cs!j] @ (drop (Suc j) (take k cs)) @ [cs!k] @ (drop (Suc k) (take (Suc i) cs))) z \ 2" using zcsk 1(2) by(simp) with A have "count_list (take (Suc i) cs) z \ 2" by auto with z_occurs_once_in_cs show "False" by auto next case (2 j) then have 1: "Suc k+j < i" by auto then have f2: "j < length (drop (Suc k) (take (Suc i) cs))" using i_in_cs by simp have 3: "(drop (Suc k) (take (Suc i) cs)) = take j (drop (Suc k) (take (Suc i) cs)) @ (drop (Suc k) (take (Suc i) cs))! j # drop (Suc j) (drop (Suc k) (take (Suc i) cs))" using f2 by(rule id_take_nth_drop) have "(drop (Suc k) (take (Suc i) cs))! j = (take (Suc i) cs) ! (Suc k+j)" apply(rule nth_drop) using i_in_cs 1 by auto also have "\ = cs ! (Suc k+j)" apply(rule nth_take) using 1 by auto finally have 4: "(drop (Suc k) (take (Suc i) cs)) = take j (drop (Suc k) (take (Suc i) cs)) @ cs! (Suc k +j) # drop (Suc j) (drop (Suc k) (take (Suc i) cs))" using 3 by auto have 5: "cs2 ! j = cs! (Suc k +j)" unfolding cs2 apply(rule nth_drop) using i_in_cs 1 by auto from 4 5 2(2) have 6: "(drop (Suc k) (take (Suc i) cs)) = take j (drop (Suc k) (take (Suc i) cs)) @ cs! k # drop (Suc j) (drop (Suc k) (take (Suc i) cs))" by auto from i_in_cs ki have 1: "k < length (take (Suc i) cs)" by auto then have 7: "(take (Suc i) cs) = (take k (take (Suc i) cs)) @ (take (Suc i) cs)!k # (drop (Suc k) (take (Suc i) cs))" by(rule id_take_nth_drop) have 9: "(take (Suc i) cs)!k = z" unfolding zcsk apply(rule nth_take) using ki by auto from 6 7 have A: "(take (Suc i) cs) = (take k (take (Suc i) cs)) @ z # take j (drop (Suc k) (take (Suc i) cs)) @ z # drop (Suc j) (drop (Suc k) (take (Suc i) cs))" using ki 9 by auto have "count_list ((take k (take (Suc i) cs)) @ z # take j (drop (Suc k) (take (Suc i) cs)) @ z # drop (Suc j) (drop (Suc k) (take (Suc i) cs))) z \ 2" by(simp) with A have "count_list (take (Suc i) cs) z \ 2" by auto with z_occurs_once_in_cs show "False" by auto qed have k_in_cs: "k < length cs" using ki i_in_cs by auto with cs1 have lenkk: "length cs1 = k+1" by auto from k_in_cs have mincsk: "min (length cs) (Suc k) = Suc k" by auto have "s_TS init h (((as@[x]@bs@[x])@cs1) @ cs2) (length (as@[x]@bs@[x])+k+1) = s_TS init h ((as@[x]@bs@[x])@(cs1)) (length (as@[x]@bs@[x])+k+1)" apply(rule s_TS_append) using cs1 cs2 k_in_cs by(simp) then have spliter: "s_TS init h ((as@[x]@bs@[x])@(cs1)) (length (as@[x]@bs@[x]@(cs1))) = s_TS init h ((as@[x]@bs@[x])@cs) (length (as@[x]@bs@[x])+k+1) " using lenkk v cs1 apply(auto) by (simp add: add.commute add.left_commute) from cs2 have "length cs2 = length cs - (Suc k)" by auto have notxbeforez: "~ x < z in s_TS init h \ (length (as @ [x] @ bs @ [x]) + k + 1)" proof (rule ccontr, goal_cases) case 1 then have a: "x < z in s_TS init h ((as@[x]@bs@[x])@(cs1)) (length (as@[x]@bs@[x]@(cs1)))" unfolding spliter assms(1) by auto have 41: "x \ set(s_TS init h ((as @ [x] @ bs @ [x]) @ cs) (length (as @ [x] @ bs @ [x]) + i))" using i_in_cs assms(4) by(simp add: s_TS_set) have 42: "z \ set(s_TS init h ((as @ [x] @ bs @ [x]) @ cs) (length (as @ [x] @ bs @ [x]) + i))" using i_in_cs zininit by(simp add: s_TS_set) have rewr: "s_TS init h ((as@[x]@bs@[x]@cs1)@cs2) (length (as@[x]@bs@[x]@cs1)+(i-k-1)) = s_TS init h (as@[x]@bs@[x]@cs) (length (as@[x]@bs@[x])+i)" using cs1 v ki apply(simp add: mincsk) by (simp add: add.commute add.left_commute) have "x < z in s_TS init h ((as@[x]@bs@[x]@cs1)@cs2) (length (as@[x]@bs@[x]@cs1)+(i-k-1))" apply(rule TS_mono) using a apply(simp) using cs2 i_in_cs ki v cs1 apply(simp) using z_lastocc zcsk apply(simp) using v assms(2) apply force using assms by(simp_all add: cs1 cs2) (* "contradiction to zmustbebeforex" *) from zmustbebeforex this[unfolded rewr ] el_n_x zcsk 41 42 not_before_in show "False" unfolding s_TS_def by fastforce qed have 1: "k < length cs" "(\j cs ! k)" "cs ! k \ x" "cs ! k \ set bs" "~ x < z in s_TS init h \ (length (as @ [x] @ bs @ [x]) + k + 1)" apply(safe) using ki i_in_cs apply(simp) using z_firstocc apply(simp) using assms(2) ki i_in_cs apply(fastforce) using z_notinbs apply(simp) using notxbeforez by auto show ?case apply(simp only: ex_nat_less_eq) apply(rule bexI[where x=k]) using 1 zcsk apply(simp) using ki by simp qed lemma nopaid: "snd (fst (TS_step_d s q)) = []" unfolding TS_step_d_def by simp lemma staysuntouched: assumes d[simp]: "distinct (fst S)" and x: "x \ set (fst S)" and y: "y \ set (fst S)" shows "set qs \ set (fst S) \ x \ set qs \ y \ set qs \ x < y in fst (config' (rTS []) S qs) = x < y in fst S" proof(induct qs rule: rev_induct) case (snoc q qs) have "x < y in fst (config' (rTS []) S (qs @ [q])) = x < y in fst (config' (rTS []) S qs)" apply(simp add: config'_snoc Step_def split_def step_def rTS_def nopaid) apply(rule xy_relativorder_mtf2) using snoc by(simp_all add: x y ) also have "\ = x < y in fst S" apply(rule snoc) using snoc by simp_all finally show ?case . qed simp lemma staysuntouched': assumes d[simp]: "distinct init" and x: "x \ set init" and y: "y \ set init" and "set qs \ set init" and "x \ set qs" and "y \ set qs" shows "x < y in fst (config (rTS []) init qs) = x < y in init" proof - let ?S="(init, fst (rTS []) init)" have "x < y in fst (config' (rTS []) ?S qs) = x < y in fst ?S" apply(rule staysuntouched) using assms by(simp_all) then show ?thesis by simp qed lemma projEmpty: "Lxy qs S = [] \ x \ S \ x \ set qs" unfolding Lxy_def by (metis filter_empty_conv) lemma Lxy_index_mono: assumes "x\S" "y\S" and "index xs x < index xs y" and "index xs y < length xs" and "x\y" shows "index (Lxy xs S) x < index (Lxy xs S) y" proof - from assms have ij: "index xs x < index xs y" and xinxs: "index xs x < length xs" and yinxs: "index xs y < length xs" by auto then have inset: "x\set xs" "y\set xs" using index_less_size_conv by fast+ from xinxs obtain a as where dec1: "a @ [xs!index xs x] @ as = xs" and a: "a = take (index xs x) xs" and "as = drop (Suc (index xs x)) xs" and length_a: "length a = index xs x" and length_as: "length as = length xs - index xs x- 1" using id_take_nth_drop by fastforce have "index xs y\length (a @ [xs!index xs x])" using length_a ij by auto then have "((a @ [xs!index xs x]) @ as) ! index xs y = as ! (index xs y-length (a @ [xs ! index xs x]))" using nth_append[where xs="a @ [xs!index xs x]" and ys="as"] by(simp) then have xsj: "xs ! index xs y = as ! (index xs y-index xs x-1)" using dec1 length_a by auto have las: "(index xs y-index xs x-1) < length as" using length_as yinxs ij by simp obtain b c where dec2: "b @ [xs!index xs y] @ c = as" and "b = take (index xs y-index xs x-1) as" "c=drop (Suc (index xs y-index xs x-1)) as" and length_b: "length b = index xs y-index xs x-1" using id_take_nth_drop[OF las] xsj by force have xs_dec: "a @ [xs!index xs x] @ b @ [xs!index xs y] @ c = xs" using dec1 dec2 by auto then have "Lxy xs S = Lxy (a @ [xs!index xs x] @ b @ [xs!index xs y] @ c) S" by(simp add: xs_dec) also have "\ = Lxy a S @ Lxy [x] S @ Lxy b S @ Lxy [y] S @ Lxy c S" by(simp add: Lxy_append Lxy_def assms inset) finally have gr: "Lxy xs S = Lxy a S @ [x] @ Lxy b S @ [y] @ Lxy c S" using assms by(simp add: Lxy_def) have "y \ set (take (index xs x) xs)" apply(rule index_take) using assms by simp then have "y \ set (Lxy (take (index xs x) xs) S )" apply(subst Lxy_set_filter) by blast with a have ynot: "y \ set (Lxy a S)" by simp have "index (Lxy xs S) y = index (Lxy a S @ [x] @ Lxy b S @ [y] @ Lxy c S) y" by(simp add: gr) also have "\ \ length (Lxy a S) + 1" using assms(5) ynot by(simp add: index_append) finally have 1: "index (Lxy xs S) y \ length (Lxy a S) + 1" . have "index (Lxy xs S) x = index (Lxy a S @ [x] @ Lxy b S @ [y] @ Lxy c S) x" by (simp add: gr) also have "\ \ length (Lxy a S)" apply(simp add: index_append) apply(subst index_less_size_conv[symmetric]) by simp finally have 2: "index (Lxy xs S) x \ length (Lxy a S)" . from 1 2 show ?thesis by linarith qed lemma proj_Cons: assumes filterd_cons: "Lxy qs S = a#as" and a_filter: "a\S" obtains pre suf where "qs = pre @ [a] @ suf" and "\x. x \ S \ x \ set pre" and "Lxy suf S = as" proof - have "set (Lxy qs S) \ set qs" using Lxy_set_filter by fast with filterd_cons have a_inq: "a \ set qs" by simp then have "index qs a < length qs" by(simp) { fix e assume eS:"e\S" assume "e\a" have "index qs a \ index qs e" proof (rule ccontr) assume "\ index qs a \ index qs e" then have 1: "index qs e < index qs a" by simp have 0: "index (Lxy qs S) a = 0" unfolding filterd_cons by simp have 2: "index (Lxy qs S) e < index (Lxy qs S) a" apply(rule Lxy_index_mono) by(fact)+ from 0 2 show "False" by linarith qed } note atfront=this let ?lastInd="index qs a" have "qs = take ?lastInd qs @ qs!?lastInd # drop (Suc ?lastInd) qs" apply(rule id_take_nth_drop) using a_inq by simp also have "\ = take ?lastInd qs @ [a] @ drop (Suc ?lastInd) qs" using a_inq by simp finally have split: "qs = take ?lastInd qs @ [a] @ drop (Suc ?lastInd) qs" . have nothingin: "\s. s\S \ s \ set (take ?lastInd qs)" apply(rule index_take) apply(case_tac "a=s") apply(simp) by (rule atfront) simp_all then have "set (Lxy (take ?lastInd qs) S) = {}" apply(subst Lxy_set_filter) by blast then have emptyPre: "Lxy (take ?lastInd qs) S = []" by blast have "a#as = Lxy qs S" using filterd_cons by simp also have "\ = Lxy (take ?lastInd qs @ [a] @ drop (Suc ?lastInd) qs) S" using split by simp also have "\ = Lxy (take ?lastInd qs) S @ (Lxy [a] S) @ Lxy (drop (Suc ?lastInd) qs) S" by(simp add: Lxy_append Lxy_def) also have "\ = a#Lxy (drop (Suc ?lastInd) qs) S" unfolding emptyPre by(simp add: Lxy_def a_filter) finally have suf: "Lxy (drop (Suc ?lastInd) qs) S = as" by simp from split nothingin suf show ?thesis .. qed lemma Lxy_rev: "rev (Lxy qs S) = Lxy (rev qs) S" apply(induct qs) by(simp_all add: Lxy_def) lemma proj_Snoc: assumes filterd_cons: "Lxy qs S = as@[a]" and a_filter: "a\S" obtains pre suf where "qs = pre @ [a] @ suf" and "\x. x \ S \ x \ set suf" and "Lxy pre S = as" proof - have "Lxy (rev qs) S = rev (Lxy qs S)" by(simp add: Lxy_rev) also have "\ = a#(rev as)" unfolding filterd_cons by simp finally have "Lxy (rev qs) S = a # (rev as)" . with a_filter obtain pre' suf' where 1: "rev qs = pre' @[a] @ suf'" and 2: "\x. x \ S \ x \ set pre'" and 3: "Lxy suf' S = rev as" using proj_Cons by metis have "qs = rev (rev qs)" by simp also have "\ = rev suf' @ [a] @ rev pre'" using 1 by simp finally have a1: "qs = rev suf' @ [a] @ rev pre'" . have "Lxy (rev suf') S = rev (Lxy suf' S)" by(simp add: Lxy_rev) also have "\ = as" using 3 by simp finally have a3: "Lxy (rev suf') S = as" . have a2: "\x. x \ S \ x \ set (rev pre')" using 2 by simp from a1 a2 a3 show ?thesis .. qed lemma sndTSconfig': "snd (config' (rTS initH) (init,[]) qs) = rev qs @ []" apply(induct qs rule: rev_induct) apply(simp add: rTS_def) by(simp add: split_def TS_step_d_def config'_snoc Step_def rTS_def) lemma projxx: fixes e a bs assumes axy: "a\{x,y}" assumes ane: "a\e" assumes exy: "e\{x,y}" assumes add: "f\{[],[e]}" assumes bsaxy: "set (bs @ [a] @ f) \ {x,y}" assumes Lxyinitxy: "Lxy init {x, y} \ {[x,y],[y,x]}" shows "a < e in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) ((bs @ [a] @ f) @ [a]))" proof - have aexy: "{a,e}={x,y}" using exy axy ane by blast let ?h="snd (Partial_Cost_Model.config' (\s. [], TS_step_d) (Lxy init {x, y}, []) (bs @ a # f))" have history: "?h = (rev f)@a#(rev bs)" using sndTSdet[of "length (bs@a#f)" "bs@a#f", unfolded rTS_def] by(simp) { fix xs s assume sinit: "s:{[a,e],[e,a]}" assume "set xs \ {a,e}" then have "fst (config' (\s. [], TS_step_d) (s, []) xs) \ {[a,e], [e,a]}" apply (induct xs rule: rev_induct) using sinit apply(simp) apply(subst config'_append2) apply(simp only: Step_def config'.simps Let_def split_def fst_conv) apply(rule stepxy) by simp_all } note staysae=this have opt: "fst (config' (\s. [], TS_step_d) (Lxy init {x, y}, []) (bs @ [a] @ f)) \ {[a,e], [e,a]}" apply(rule staysae) using Lxyinitxy exy axy ane apply fast unfolding aexy by(fact bsaxy) have contr: " (\x. 0 < (if e = x then 0 else index [a] x + 1)) = False" proof (rule ccontr, goal_cases) case 1 then have "\x. 0 < (if e = x then 0 else index [a] x + 1)" by simp then have "0 < (if e = e then 0 else index [a] e + 1)" by blast then have "0<0" by simp then show "False" by auto qed show "a < e in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) ((bs @ [a] @ f) @ [a]))" apply(subst config_append) apply(simp add: rTS_def Step_def split_def) apply(subst TS_step_d_def) apply(simp only: history) using opt ane add apply(auto simp: step_def) apply(simp add: before_in_def) apply(simp add: before_in_def) apply(simp add: before_in_def contr) apply(simp add: mtf2_def swap_def before_in_def) apply(auto simp add: before_in_def contr) apply (metis One_nat_def add_is_1 count_list.simps(1) le_Suc_eq) by(simp add: mtf2_def swap_def) qed lemma oneposs: assumes "set xs = {x,y}" assumes "x\y" assumes "distinct xs" assumes True: "x index xs y = 1" using len2 by linarith have "xs = take 1 xs @ xs!1 # drop (Suc 1) xs" apply(rule id_take_nth_drop) using len2 by simp also have "\ = take 1 xs @ [xs!1]" using len2 by simp also have "take 1 xs = take 0 (take 1 xs) @ (take 1 xs)!0 # drop (Suc 0) (take 1 xs)" apply(rule id_take_nth_drop) using len2 by simp also have "\ = [xs!0]" by(simp) finally have "xs = [xs!0, xs!1]" by simp also have "\ = [xs!(index xs x), xs!index xs y]" using f by simp also have "\ = [x,y]" using assms by(simp) finally show "xs = [x,y]" . qed lemma twoposs: assumes "set xs = {x,y}" assumes "x\y" assumes "distinct xs" shows "xs \ {[x,y], [y,x]}" proof (cases "x index xs y = 1" using len2 by linarith have "xs = take 1 xs @ xs!1 # drop (Suc 1) xs" apply(rule id_take_nth_drop) using len2 by simp also have "\ = take 1 xs @ [xs!1]" using len2 by simp also have "take 1 xs = take 0 (take 1 xs) @ (take 1 xs)!0 # drop (Suc 0) (take 1 xs)" apply(rule id_take_nth_drop) using len2 by simp also have "\ = [xs!0]" by(simp) finally have "xs = [xs!0, xs!1]" by simp also have "\ = [xs!(index xs x), xs!index xs y]" using f by simp also have "\ = [x,y]" using assms by(simp) finally have "xs = [x,y]" . then show ?thesis by simp next case False from assms have len2: "length xs = 2" using distinct_card[OF assms(3)] by fastforce from False have "y index xs x = 1" using len2 by linarith have "xs = take 1 xs @ xs!1 # drop (Suc 1) xs" apply(rule id_take_nth_drop) using len2 by simp also have "\ = take 1 xs @ [xs!1]" using len2 by simp also have "take 1 xs = take 0 (take 1 xs) @ (take 1 xs)!0 # drop (Suc 0) (take 1 xs)" apply(rule id_take_nth_drop) using len2 by simp also have "\ = [xs!0]" by(simp) finally have "xs = [xs!0, xs!1]" by simp also have "\ = [xs!(index xs y), xs!index xs x]" using f by simp also have "\ = [y,x]" using assms by(simp) finally have "xs = [y,x]" . then show ?thesis by simp qed lemma TS_pairwise': assumes "qs \ {xs. set xs \ set init}" "(x, y) \ {(x, y). x \ set init \ y \ set init \ x \ y}" "x \ y" "distinct init" shows "Pbefore_in x y (embed (rTS [])) qs init = Pbefore_in x y (embed (rTS [])) (Lxy qs {x, y}) (Lxy init {x, y})" proof - from assms have xyininit: "{x, y} \ set init" and qsininit: "set qs \ set init" by auto note dinit=assms(4) from assms have xny: "x\y" by simp have Lxyinitxy: "Lxy init {x, y} \ {[x, y], [y, x]}" apply(rule twoposs) apply(subst Lxy_set_filter) using xyininit apply fast using xny Lxy_distinct[OF dinit] by simp_all have lq_s: "set (Lxy qs {x, y}) \ {x,y}" by (simp add: Lxy_set_filter) (* projected history *) let ?pH = "snd (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" have "?pH =snd (TSdet (Lxy init {x, y}) [] (Lxy qs {x, y}) (length (Lxy qs {x, y})))" by(simp) also have "\ = rev (take (length (Lxy qs {x, y})) (Lxy qs {x, y})) @ []" apply(rule sndTSdet) by simp finally have pH: "?pH = rev (Lxy qs {x, y})" by simp let ?pQs = "(Lxy qs {x, y})" have A: " x < y in fst (config\<^sub>p (rTS []) init qs) = x < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" proof(cases "?pQs" rule: rev_cases) case Nil then have xqs: "x \ set qs" and yqs: "y \ set qs" by(simp_all add: projEmpty) have " x < y in fst (config\<^sub>p (rTS []) init qs) = x < y in init" apply(rule staysuntouched') using assms xqs yqs by(simp_all) also have "\ = x < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" unfolding Nil apply(simp) apply(rule Lxy_mono) using xyininit dinit by(simp_all) finally show ?thesis . next case (snoc as a) then have "a\set (Lxy qs {x, y})" by (simp) then have axy: "a\{x,y}" by(simp add: Lxy_set_filter) with xyininit have ainit: "a\set init" by auto note a=snoc from a axy obtain pre suf where qs: "qs = pre @ [a] @ suf" and nosuf: "\e. e \ {x,y} \ e \ set suf" and pre: "Lxy pre {x,y} = as" using proj_Snoc by metis show ?thesis proof (cases "as" rule: rev_cases) case Nil from pre Nil have xqs: "x \ set pre" and yqs: "y \ set pre" by(simp_all add: projEmpty) from xqs yqs axy have "a \ set pre" by blast then have noocc: "index (rev pre) a = length (rev pre)" by simp have " x < y in fst (config\<^sub>p (rTS []) init qs) = x < y in fst (config\<^sub>p (rTS []) init ((pre @ [a]) @ suf))" by(simp add: qs) also have "\ = x < y in fst (config\<^sub>p (rTS []) init (pre @ [a]))" apply(subst config_append) apply(rule staysuntouched) using assms xqs yqs qs nosuf by(simp_all) also have "\ = x < y in fst (config\<^sub>p (rTS []) init pre)" apply(subst config_append) apply(simp add: rTS_def Step_def split_def) apply(simp only: TS_step_d_def) apply(simp only: sndTSconfig'[unfolded rTS_def]) by(simp add: noocc step_def) also have "\ = x < y in init" apply(rule staysuntouched') using assms xqs yqs qs by(simp_all) also have "\ = x < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" unfolding a Nil apply(simp add: Step_def split_def rTS_def TS_step_d_def step_def) apply(rule Lxy_mono) using xyininit dinit by(simp_all) finally show ?thesis . next case (snoc bs b) note b=this with a have "b\set (Lxy qs {x, y})" by (simp) then have bxy: "b\{x,y}" by(simp add: Lxy_set_filter) with xyininit have binit: "b\set init" by auto from b pre have "Lxy pre {x,y} = bs @ [b]" by simp with bxy obtain pre2 suf2 where bs: "pre = pre2 @ [b] @ suf2" and nosuf2: "\e. e \ {x,y} \ e \ set suf2" and pre2: "Lxy pre2 {x,y} = bs" using proj_Snoc by metis from bs qs have qs2: "qs = pre2 @ [b] @ suf2 @ [a] @ suf" by simp show ?thesis proof (cases "a=b") case True note ab=this let ?qs ="(pre2 @ [a] @ suf2 @ [a]) @ suf" { fix e assume ane: "a\e" assume exy: "e\{x,y}" have "a < e in fst (config\<^sub>p (rTS []) init qs) = a < e in fst (config\<^sub>p (rTS []) init ?qs)" using True qs2 by(simp) also have "\ = a < e in fst (config\<^sub>p (rTS []) init (pre2 @ [a] @ suf2 @ [a]))" apply(subst config_append) apply(rule staysuntouched) using assms qs nosuf apply(simp_all) using exy xyininit apply fast using nosuf axy apply(simp) using nosuf exy by simp also have "\" apply(simp) apply(rule twotox[unfolded s_TS_def, simplified]) using nosuf2 exy apply(simp) using assms apply(simp_all) using axy xyininit apply fast using exy xyininit apply fast using nosuf2 axy apply(simp) using ane by simp finally have "a < e in fst (config\<^sub>p (rTS []) init qs)" by simp } note full=this have "set (bs @ [a]) \ set (Lxy qs {x, y})" using a b by auto also have "\ = {x,y} \ set qs" by (rule Lxy_set_filter) also have "\ \ {x,y}" by simp finally have bsaxy: "set (bs @ [a]) \ {x,y}" . with xny show ?thesis proof(cases "x=a") case True have 1: "a < y in fst (config\<^sub>p (rTS []) init qs)" apply(rule full) using True xny apply blast by simp have "a < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})) = a < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) ((bs @ [a] @ []) @ [a]))" using a b ab by simp also have "\" apply(rule projxx[where bs=bs and f="[]"]) using True apply blast using a b True ab xny Lxyinitxy bsaxy by(simp_all) finally show ?thesis using True 1 by simp next case False with axy have ay: "a=y" by blast have 1: "a < x in fst (config\<^sub>p (rTS []) init qs)" apply(rule full) using False xny apply blast by simp have "a < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})) = a < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) ((bs @ [a] @ []) @ [a]))" using a b ab by simp also have "\" apply(rule projxx[where bs=bs and f="[]"]) using True axy apply blast using a b True ab xny Lxyinitxy ay bsaxy by(simp_all) finally have 2: "a < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" . have "x < y in fst (config\<^sub>p (rTS []) init qs) = (\ y < x in fst (config\<^sub>p (rTS []) init qs))" apply(subst not_before_in) using assms by(simp_all) also have "\ = False" using 1 ay by simp also have "\ = (\ y < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})))" using 2 ay by simp also have "\ = x < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" apply(subst not_before_in) using assms by(simp_all add: Lxy_set_filter) finally show ?thesis . qed next case False note ab=this show ?thesis proof (cases "bs" rule: rev_cases) case Nil with a b have "Lxy qs {x, y} = [b,a]" by simp from pre2 Nil have xqs: "x \ set pre2" and yqs: "y \ set pre2" by(simp_all add: projEmpty) from xqs yqs bxy have "b \ set pre2" by blast then have noocc2: "index (rev pre2) b = length (rev pre2)" by simp from axy nosuf2 have "a \ set suf2" by blast with xqs yqs axy False have "a \ set ((pre2 @ b # suf2))" by(auto) then have noocc: "index (rev (pre2 @ b # suf2) @ []) a = length (rev (pre2 @ b # suf2))" by simp have " x < y in fst (config\<^sub>p (rTS []) init qs) = x < y in fst (config\<^sub>p (rTS []) init ((((pre2 @ [b]) @ suf2) @ [a]) @ suf))" by(simp add: qs2) also have "\ = x < y in fst (config\<^sub>p (rTS []) init (((pre2 @ [b]) @ suf2) @ [a]))" apply(subst config_append) apply(rule staysuntouched) using assms xqs yqs qs nosuf by(simp_all) also have "\ = x < y in fst (config\<^sub>p (rTS []) init ((pre2 @ [b]) @ suf2))" apply(subst config_append) apply(simp add: rTS_def Step_def split_def) apply(simp only: TS_step_d_def) apply(simp only: sndTSconfig'[unfolded rTS_def]) apply(simp only: noocc) by (simp add: step_def) also have "\ = x < y in fst (config\<^sub>p (rTS []) init (pre2 @ [b]))" apply(subst config_append) apply(rule staysuntouched) using assms xqs yqs qs2 nosuf2 by(simp_all) also have "\ = x < y in fst (config\<^sub>p (rTS []) init (pre2))" apply(subst config_append) apply(simp add: rTS_def Step_def split_def) apply(simp only: TS_step_d_def) apply(simp only: sndTSconfig'[unfolded rTS_def]) by(simp add: noocc2 step_def) also have "\ = x < y in init" apply(rule staysuntouched') using assms xqs yqs qs2 by(simp_all) also have "\ = x < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" unfolding a b Nil using False apply(simp add: Step_def split_def rTS_def TS_step_d_def step_def) apply(rule Lxy_mono) using xyininit dinit by(simp_all) finally show ?thesis . next case (snoc cs c) note c=this with a b have "c\set (Lxy qs {x, y})" by (simp) then have cxy: "c\{x,y}" by(simp add: Lxy_set_filter) from c pre2 have "Lxy pre2 {x,y} = cs @ [c]" by simp with cxy obtain pre3 suf3 where cs: "pre2 = pre3 @ [c] @ suf3" and nosuf3: "\e. e \ {x,y} \ e \ set suf3" and pre3: "Lxy pre3 {x,y} = cs" using proj_Snoc by metis let ?qs=" pre3 @ [c] @ suf3 @ [b] @ suf2 @ [a] @ suf" from bs cs qs have qs2: "qs = ?qs" by simp show ?thesis proof(cases "c=a") case True (* aba *) note ca=this have "a < b in fst (config\<^sub>p (rTS []) init qs) = a < b in fst (config\<^sub>p (rTS []) init ((pre3 @ a # (suf3 @ [b] @ suf2) @ [a]) @ suf))" using qs2 True by simp also have "\ = a < b in fst (config\<^sub>p (rTS []) init (pre3 @ a # (suf3 @ [b] @ suf2) @ [a]))" apply(subst config_append) apply(rule staysuntouched) using assms qs nosuf apply(simp_all) using bxy xyininit apply(fast) using nosuf axy bxy by(simp_all) also have "..." apply(rule twotox[unfolded s_TS_def, simplified]) using nosuf2 nosuf3 bxy apply(simp) using assms apply(simp_all) using axy xyininit apply(fast) using bxy xyininit apply(fast) using ab nosuf2 nosuf3 axy apply(simp) using ab by simp finally have full: "a < b in fst (config\<^sub>p (rTS []) init qs)" by simp have "set (cs @ [a] @ [b]) \ set (Lxy qs {x, y})" using a b c by auto also have "\ = {x,y} \ set qs" by (rule Lxy_set_filter) also have "\ \ {x,y}" by simp finally have csabxy: "set (cs @ [a] @ [b]) \ {x,y}" . with xny show ?thesis proof(cases "x=a") case True with xny ab bxy have bisy: "b=y" by blast have 1: "x < y in fst (config\<^sub>p (rTS []) init qs)" using full True bisy by simp have "a < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})) = a < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) ((cs @ [a] @ [b]) @ [a]))" using a b c ca ab by simp also have "\" apply(rule projxx) using True apply blast using a b True ab xny Lxyinitxy csabxy by(simp_all) finally show ?thesis using 1 True by simp next case False with axy have ay: "a=y" by blast with xny ab bxy have bisx: "b=x" by blast have 1: "y < x in fst (config\<^sub>p (rTS []) init qs)" using full ay bisx by simp have "a < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})) = a < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) ((cs @ [a] @ [b]) @ [a]))" using a b c ca ab by simp also have "\" apply(rule projxx) using a b True ab xny Lxyinitxy csabxy False by(simp_all) finally have 2: "a < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" . have "x < y in fst (config\<^sub>p (rTS []) init qs) = (\ y < x in fst (config\<^sub>p (rTS []) init qs))" apply(subst not_before_in) using assms by(simp_all) also have "\ = False" using 1 ay by simp also have "\ = (\ y < x in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})))" using 2 ay by simp also have "\ = x < y in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" apply(subst not_before_in) using assms by(simp_all add: Lxy_set_filter) finally show ?thesis . qed next case False (* bba *) then have cb: "c=b" using bxy cxy axy ab by blast let ?cs = "suf2 @ [a] @ suf" let ?i = "index ?cs a" have aed: "(\j a)" by (metis add.right_neutral axy index_Cons index_append nosuf2 nth_append nth_mem) have "?i < length ?cs \ (\j ?cs ! ?i) \ ?cs ! ?i \ b \ ?cs ! ?i \ set suf3 \ b < ?cs ! ?i in s_TS init [] qs (length (pre3 @ [b] @ suf3 @ [b]) + ?i + 1)" apply(rule casexxy) using cb qs2 apply(simp) using bxy ab nosuf2 nosuf apply(simp) using bs qs qsininit apply(simp) using bxy xyininit apply(blast) apply(fact) using nosuf3 bxy apply(simp) using cs bs qs qsininit by(simp_all) then have inner: "b < a in s_TS init [] qs (length (pre3 @ [b] @ suf3 @ [b]) + ?i + 1)" using ab nosuf3 axy bxy aed by(simp) let ?n = "(length (pre3 @ [b] @ suf3 @ [b]) + ?i + 1)" let ?inner="(config\<^sub>p (rTS []) init (take (length (pre3 @ [b] @ suf3 @ [b]) + ?i + 1) ?qs))" have "b < a in fst (config\<^sub>p (rTS []) init qs) = b < a in fst (config\<^sub>p (rTS []) init (take ?n ?qs @ drop ?n ?qs))" using qs2 by simp also have "\ = b < a in fst (config' (rTS []) ?inner suf)" apply(simp only: config_append drop_append) using nosuf2 axy by(simp add: index_append config_append) also have "\ = b < a in fst ?inner" apply(rule staysuntouched) using assms bxy xyininit qs nosuf apply(simp_all) using bxy xyininit apply(blast) using axy xyininit by (blast) also have "\ = True" using inner by(simp add: s_TS_def qs2) finally have full: "b < a in fst (config\<^sub>p (rTS []) init qs)" by simp have "set (cs @ [b] @ []) \ set (Lxy qs {x, y})" using a b c by auto also have "\ = {x,y} \ set qs" by (rule Lxy_set_filter) also have "\ \ {x,y}" by simp finally have csbxy: "set (cs @ [b] @ []) \ {x,y}" . have "set (Lxy init {x, y}) = {x,y} \ set init" by(rule Lxy_set_filter) also have "\ = {x,y}" using xyininit by fast also have "\ = {b,a}" using axy bxy ab by fast finally have r: "set (Lxy init {x, y}) = {b, a}" . let ?confbef="(config\<^sub>p (rTS []) (Lxy init {x, y}) ((cs @ [b] @ []) @ [b]))" have f1: "b < a in fst ?confbef" apply(rule projxx) using bxy ab axy a b c csbxy Lxyinitxy by(simp_all) have 1: "fst ?confbef = [b,a]" apply(rule oneposs) using ab axy bxy xyininit Lxy_distinct[OF dinit] f1 r by(simp_all) have 2: "snd (Partial_Cost_Model.config' (\s. [], TS_step_d) (Lxy init {x, y}, []) (cs @ [b, b])) = [b,b]@(rev cs)" using sndTSdet[of "length (cs @ [b, b])" "(cs @ [b, b])", unfolded rTS_def] by(simp) have "b < a in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})) = b < a in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (((cs @ [b] @ []) @ [b])@[a]))" using a b c cb by(simp) also have "\" apply(subst config_append) using 1 2 ab apply(simp add: step_def Step_def split_def rTS_def TS_step_d_def) by(simp add: before_in_def) finally have projected: "b < a in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" . have 1: "{x,y} = {a,b}" using ab axy bxy by fast with xny show ?thesis proof(cases "x=a") case True with 1 xny have y: "y=b" by fast have "a < b in fst (config\<^sub>p (rTS []) init qs) = (\ b < a in fst (config\<^sub>p (rTS []) init qs))" apply(subst not_before_in) using binit ainit ab by(simp_all) also have "\ = False" using full by simp also have "\ = (\ b < a in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y})))" using projected by simp also have "\ = a < b in fst (config\<^sub>p (rTS []) (Lxy init {x, y}) (Lxy qs {x, y}))" apply(subst not_before_in) using binit ainit ab axy bxy by(simp_all add: Lxy_set_filter) finally show ?thesis using True y by simp next case False with 1 xny have y: "y=a" "x=b" by fast+ with full projected show ?thesis by fast qed qed (* end of (c=a) *) qed (* end of snoc cs c *) qed (* end of (a=b) *) qed (* end snoc bs b *) qed (* end snoc as a *) show ?thesis unfolding Pbefore_in_def apply(subst config_embed) apply(subst config_embed) apply(simp) by (rule A) qed theorem TS_pairwise: "pairwise (embed (rTS []))" apply(rule pairwise_property_lemma) apply(rule TS_pairwise') by (simp_all add: rTS_def TS_step_d_def) subsection "TS is 2-compet" lemma TS_compet': "pairwise (embed (rTS [])) \ \s0\{init::(nat list). distinct init \ init\[]}. \b\0. \qs\{x. set x \ set s0}. T\<^sub>p_on_rand (embed (rTS [])) s0 qs \ (2::real) * T\<^sub>p_opt s0 qs + b" unfolding rTS_def proof (rule factoringlemma_withconstant, goal_cases) case 5 show ?case proof (safe, goal_cases) case (1 init) note out=this show ?case apply(rule exI[where x=2]) apply(simp) proof (safe, goal_cases) case (1 qs a b) then have a: "a\b" by simp have twist: "{a,b}={b, a}" by auto have b1: "set (Lxy qs {a, b}) \ {a, b}" unfolding Lxy_def by auto with this[unfolded twist] have b2: "set (Lxy qs {b, a}) \ {b, a}" by(auto) have "set (Lxy init {a, b}) = {a,b} \ (set init)" apply(induct init) unfolding Lxy_def by(auto) with 1 have A: "set (Lxy init {a, b}) = {a,b}" by auto have "finite {a,b}" by auto from out have B: "distinct (Lxy init {a, b})" unfolding Lxy_def by auto have C: "length (Lxy init {a, b}) = 2" using distinct_card[OF B, unfolded A] using a by auto have "{xs. set xs = {a,b} \ distinct xs \ length xs =(2::nat)} = { [a,b], [b,a] }" apply(auto simp: a a[symmetric]) proof (goal_cases) case (1 xs) from 1(4) obtain x xs' where r:"xs=x#xs'" by (metis Suc_length_conv add_2_eq_Suc' append_Nil length_append) with 1(4) have "length xs' = 1" by auto then obtain y where s: "[y] = xs'" by (metis One_nat_def length_0_conv length_Suc_conv) from r s have t: "[x,y] = xs" by auto moreover from t 1(1) have "x=b" using doubleton_eq_iff 1(2) by fastforce moreover from t 1(1) have "y=a" using doubleton_eq_iff 1(2) by fastforce ultimately show ?case by auto qed with A B C have pos: "(Lxy init {a, b}) = [a,b] \ (Lxy init {a, b}) = [b,a]" by auto { fix a::nat fix b::nat fix qs assume as: "a \ b" "set qs \ {a, b}" have "T_on_rand' (embed (rTS [])) (fst (embed (rTS [])) [a,b] \ (\is. return_pmf ([a,b], is))) qs = T\<^sub>p_on (rTS []) [a, b] qs" by (rule T_on_embed[symmetric]) also from as have "\ \ 2 * T\<^sub>p_opt [a, b] qs + 2" using TS_OPT2' by fastforce finally have "T_on_rand' (embed (rTS [])) (fst (embed (rTS [])) [a,b] \ (\is. return_pmf ([a,b], is))) qs \ 2 * T\<^sub>p_opt [a, b] qs + 2" . } note ye=this show ?case apply(cases "(Lxy init {a, b}) = [a,b]") using ye[OF a b1, unfolded rTS_def] apply(simp) using pos ye[OF a[symmetric] b2, unfolded rTS_def] by(simp add: twist) qed qed next case 6 show ?case unfolding TS_step_d_def by (simp add: split_def TS_step_d_def) next case (7 init qs x) then show ?case apply(induct x) by (simp_all add: rTS_def split_def take_Suc_conv_app_nth config'_rand_snoc ) next case 4 then show ?case by simp qed (simp_all) lemma TS_compet: "compet_rand (embed (rTS [])) 2 {init. distinct init \ init \ []}" unfolding compet_rand_def static_def using TS_compet'[OF TS_pairwise] by simp end