theory Scratch imports Main begin lemma "(\f :: nat \ nat. mono f \ f ` {0 ..< size xs} = {0 ..< size ys} \ (\i < size xs. xs!i = ys! (f i)) \ (\i. i + 1 < size xs \ (xs!i = xs!(i+1) \ f i = f(i+1)))) \ remdups_adj xs = ys" proof (induct xs arbitrary: ys rule: remdups_adj.induct) case 1 then show ?case by auto next case (2 x) then obtain f where f_img: "f ` {0 ..< size [x]} = {0 ..< size ys}" and f_nth: "\i. i < size [x] \ [x]!i = ys!(f i)" by metis have "length ys = card (f ` {0 ..< size [x]})" using f_img by auto then have "length ys = 1" by auto moreover then have "f 0 = 0" using f_img by auto ultimately show ?case using f_nth by (cases ys) auto next case (3 x1 x2 xs) from "3.prems" obtain f where f_mono: "mono f" and f_img: "f ` {0..i. i < length (x1 # x2 # xs) \ (x1 # x2 # xs) ! i = ys ! f i" and f_nth2: "\i. i + 1 < length (x1 # x2 #xs) \ ((x1 # x2 # xs) ! i = (x1 # x2 # xs) ! (i + 1)) = (f i = f (i + 1))" by metis show ?case proof cases assume "x1 = x2" let ?f' = "f o Suc" have "remdups_adj (x1 # xs) = ys" proof (intro "3.hyps" exI conjI impI allI) show "x1 = x2" by fact next show "mono ?f'" using f_mono by (simp add: mono_iff_le_Suc) next have "?f' ` {0 ..< length (x1 # xs)} = f ` {Suc 0 ..< length (x1 # x2 # xs)}" by safe (fastforce, rename_tac x, case_tac x, auto) also have "\ = f ` {0 ..< length (x1 # x2 # xs)}" proof - have "f 0 = f (Suc 0)" using \x1 = x2\ f_nth2[of 0] by simp then show ?thesis by safe (fastforce, rename_tac x, case_tac x, auto) qed also have "\ = {0 ..< length ys}" by fact finally show "?f' ` {0 ..< length (x1 # xs)} = {0 ..< length ys}" . next fix i assume "i < length (x1 # xs)" then show "(x1 # xs) ! i = ys ! ?f' i" using f_nth1[of "Suc i"] \x1 = x2\ by simp next fix i assume "i + 1 < length (x1 # xs)" then show "((x1 # xs) ! i = (x1 # xs) ! (i + 1)) = (?f' i = ?f' (i + 1))" using f_nth2[of "Suc i"] \x1 = x2\ by simp_all qed then show ?thesis using \x1 = x2\ by simp next assume "x1 \ x2" then have "f 0 \ f (Suc 0)" using f_nth2[of 0] by auto have "2 \ length ys" proof - have "2 = card {f 0, f 1}" using \f 0 \ _\ by simp also have "\ \ card (f ` ({0, 1} \ {2..< Suc (Suc (length xs))}))" (is "_ \ card (f ` ?S)") by (rule card_mono) auto also have "?S = {0..) = length ys" using f_img by simp finally show "2 \ length ys" . qed have "f 0 = 0" proof (rule ccontr) assume "f 0 \ 0" then have "\i. 0 \ f i" using f_mono by (metis le0 le_0_eq monoD) then have "0 \ f ` {0 ..< length (x1 # x2 # xs)}" by (auto simp del: neq0_conv) then show False using f_img by auto qed have "f (Suc 0) = Suc 0" proof (rule ccontr) assume "f (Suc 0) \ Suc 0" then have "Suc 0 < f (Suc 0)" using f_nth2[of 0] \x1 \ x2\ \f 0 = 0\ by auto then have "\i. Suc 0 < f (Suc i)" using f_mono by (metis Suc_le_mono le0 less_le_trans monoD) then have "\i. Suc 0 \ f i" using \f 0 = 0\ by (metis less_irrefl_nat not0_implies_Suc) then have "Suc 0 \ f ` {0 ..< length (x1 # x2 # xs)}" by auto then show False using f_img \2 \ length ys\ by auto qed obtain ys' where "ys = x1 # x2 # ys'" using \2 \ length ys\ f_nth1[of 0] f_nth1[of 1] apply (cases ys) apply (rename_tac [2] ys', case_tac [2] ys') by (auto simp: \f 0 = 0\ \f (Suc 0) = Suc 0\) def f' \ "\x. f (Suc x) - 1" { fix i have "Suc 0 \ f (Suc 0)" using f_nth2[of 0] \x1 \ x2\ \f 0 = 0\ by auto also have "\ \ f (Suc i)" using f_mono by (rule monoD) arith finally have "Suc 0 \ f (Suc i)" . } note X = this { fix i have "f (Suc i) = Suc (f' i)" using X[of i] by (auto simp: f'_def) } note f_Suc = this have "remdups_adj (x2 # xs) = (x2 # ys')" proof (intro "3.hyps" exI conjI impI allI) show "x1 \ x2" by fact next show "mono f'" using X f_mono by (auto simp: f'_def mono_iff_le_Suc le_diff_iff) next have "f' ` {0 ..< length (x2 # xs)} = (\x. f x - 1) ` {0 ..< length (x1 # x2 #xs)}" apply safe apply (rename_tac [!] n, case_tac [!] n) apply (auto simp: f'_def \f 0 = 0\ \f (Suc 0) = Suc 0\ intro: rev_image_eqI) done also have "\ = (\x. x - 1) ` f ` {0 ..< length (x1 # x2 #xs)}" by (auto simp: image_comp) also have "\ = (\x. x - 1) ` {0 ..< length ys}" by (simp only: f_img) also have "\ = {0 ..< length (x2 # ys')}" using \ys = _\ by (fastforce intro: rev_image_eqI) finally show "f' ` {0 ..< length (x2 # xs)} = {0 ..< length (x2 # ys')}" . next fix i assume "i < length (x2 # xs)" then show "(x2 # xs) ! i = (x2 # ys') ! f' i" using f_nth1[of "Suc i"] \x1 \ x2\ f_Suc by (simp add: \ys = _\) next fix i assume "i + 1 < length (x2 # xs)" then show "((x2 # xs) ! i = (x2 # xs) ! (i + 1)) = (f' i = f' (i + 1))" using f_nth2[of "Suc i"] \x1 \ x2\ f_Suc by (auto simp: \ys = _\) qed then show ?case using \ys = _\ \x1 \ x2\ by simp qed qed end