| author | wenzelm | 
| Tue, 02 Jul 2024 23:29:46 +0200 | |
| changeset 80482 | 2136ecf06a4c | 
| parent 78653 | 7ed1759fe1bd | 
| permissions | -rw-r--r-- | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 1 | (* Author: Tobias Nipkow, Daniel Stüwe *) | 
| 61784 | 2 | |
| 62130 | 3 | section \<open>1-2 Brother Tree Implementation of Sets\<close> | 
| 61784 | 4 | |
| 5 | theory Brother12_Set | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 6 | imports | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 7 | Cmp | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 8 | Set_Specs | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 9 | "HOL-Number_Theory.Fib" | 
| 61784 | 10 | begin | 
| 11 | ||
| 12 | subsection \<open>Data Type and Operations\<close> | |
| 13 | ||
| 14 | datatype 'a bro = | |
| 15 | N0 | | |
| 16 | N1 "'a bro" | | |
| 17 | N2 "'a bro" 'a "'a bro" | | |
| 18 | (* auxiliary constructors: *) | |
| 19 | L2 'a | | |
| 20 | N3 "'a bro" 'a "'a bro" 'a "'a bro" | |
| 21 | ||
| 68431 | 22 | definition empty :: "'a bro" where | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 23 | "empty = N0" | 
| 68431 | 24 | |
| 61784 | 25 | fun inorder :: "'a bro \<Rightarrow> 'a list" where | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 26 | "inorder N0 = []" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 27 | "inorder (N1 t) = inorder t" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 28 | "inorder (N2 l a r) = inorder l @ a # inorder r" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 29 | "inorder (L2 a) = [a]" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 30 | "inorder (N3 t1 a1 t2 a2 t3) = inorder t1 @ a1 # inorder t2 @ a2 # inorder t3" | 
| 61784 | 31 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 32 | fun isin :: "'a bro \<Rightarrow> 'a::linorder \<Rightarrow> bool" where | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 33 | "isin N0 x = False" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 34 | "isin (N1 t) x = isin t x" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 35 | "isin (N2 l a r) x = | 
| 61784 | 36 | (case cmp x a of | 
| 37 | LT \<Rightarrow> isin l x | | |
| 38 | EQ \<Rightarrow> True | | |
| 39 | GT \<Rightarrow> isin r x)" | |
| 40 | ||
| 41 | fun n1 :: "'a bro \<Rightarrow> 'a bro" where | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 42 | "n1 (L2 a) = N2 N0 a N0" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 43 | "n1 (N3 t1 a1 t2 a2 t3) = N2 (N2 t1 a1 t2) a2 (N1 t3)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 44 | "n1 t = N1 t" | 
| 61784 | 45 | |
| 46 | hide_const (open) insert | |
| 47 | ||
| 48 | locale insert | |
| 49 | begin | |
| 50 | ||
| 51 | fun n2 :: "'a bro \<Rightarrow> 'a \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 52 | "n2 (L2 a1) a2 t = N3 N0 a1 N0 a2 t" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 53 | "n2 (N3 t1 a1 t2 a2 t3) a3 (N1 t4) = N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 54 | "n2 (N3 t1 a1 t2 a2 t3) a3 t4 = N3 (N2 t1 a1 t2) a2 (N1 t3) a3 t4" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 55 | "n2 t1 a1 (L2 a2) = N3 t1 a1 N0 a2 N0" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 56 | "n2 (N1 t1) a1 (N3 t2 a2 t3 a3 t4) = N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 57 | "n2 t1 a1 (N3 t2 a2 t3 a3 t4) = N3 t1 a1 (N1 t2) a2 (N2 t3 a3 t4)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 58 | "n2 t1 a t2 = N2 t1 a t2" | 
| 61784 | 59 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 60 | fun ins :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 61 | "ins x N0 = L2 x" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 62 | "ins x (N1 t) = n1 (ins x t)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 63 | "ins x (N2 l a r) = | 
| 61789 | 64 | (case cmp x a of | 
| 65 | LT \<Rightarrow> n2 (ins x l) a r | | |
| 66 | EQ \<Rightarrow> N2 l a r | | |
| 67 | GT \<Rightarrow> n2 l a (ins x r))" | |
| 61784 | 68 | |
| 69 | fun tree :: "'a bro \<Rightarrow> 'a bro" where | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 70 | "tree (L2 a) = N2 N0 a N0" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 71 | "tree (N3 t1 a1 t2 a2 t3) = N2 (N2 t1 a1 t2) a2 (N1 t3)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 72 | "tree t = t" | 
| 61784 | 73 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 74 | definition insert :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 75 | "insert x t = tree(ins x t)" | 
| 61784 | 76 | |
| 77 | end | |
| 78 | ||
| 79 | locale delete | |
| 80 | begin | |
| 81 | ||
| 82 | fun n2 :: "'a bro \<Rightarrow> 'a \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 83 | "n2 (N1 t1) a1 (N1 t2) = N1 (N2 t1 a1 t2)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 84 | "n2 (N1 (N1 t1)) a1 (N2 (N1 t2) a2 (N2 t3 a3 t4)) = | 
| 61784 | 85 | N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" | | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 86 | "n2 (N1 (N1 t1)) a1 (N2 (N2 t2 a2 t3) a3 (N1 t4)) = | 
| 61784 | 87 | N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" | | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 88 | "n2 (N1 (N1 t1)) a1 (N2 (N2 t2 a2 t3) a3 (N2 t4 a4 t5)) = | 
| 61784 | 89 | N2 (N2 (N1 t1) a1 (N2 t2 a2 t3)) a3 (N1 (N2 t4 a4 t5))" | | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 90 | "n2 (N2 (N1 t1) a1 (N2 t2 a2 t3)) a3 (N1 (N1 t4)) = | 
| 61784 | 91 | N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" | | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 92 | "n2 (N2 (N2 t1 a1 t2) a2 (N1 t3)) a3 (N1 (N1 t4)) = | 
| 61784 | 93 | N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" | | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 94 | "n2 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)) a5 (N1 (N1 t5)) = | 
| 61784 | 95 | N2 (N1 (N2 t1 a1 t2)) a2 (N2 (N2 t3 a3 t4) a5 (N1 t5))" | | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 96 | "n2 t1 a1 t2 = N2 t1 a1 t2" | 
| 61784 | 97 | |
| 68020 | 98 | fun split_min :: "'a bro \<Rightarrow> ('a \<times> 'a bro) option" where
 | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 99 | "split_min N0 = None" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 100 | "split_min (N1 t) = | 
| 68020 | 101 | (case split_min t of | 
| 61784 | 102 | None \<Rightarrow> None | | 
| 103 | Some (a, t') \<Rightarrow> Some (a, N1 t'))" | | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 104 | "split_min (N2 t1 a t2) = | 
| 68020 | 105 | (case split_min t1 of | 
| 61784 | 106 | None \<Rightarrow> Some (a, N1 t2) | | 
| 107 | Some (b, t1') \<Rightarrow> Some (b, n2 t1' a t2))" | |
| 108 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 109 | fun del :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 110 | "del _ N0 = N0" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 111 | "del x (N1 t) = N1 (del x t)" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 112 | "del x (N2 l a r) = | 
| 61784 | 113 | (case cmp x a of | 
| 114 | LT \<Rightarrow> n2 (del x l) a r | | |
| 115 | GT \<Rightarrow> n2 l a (del x r) | | |
| 68020 | 116 | EQ \<Rightarrow> (case split_min r of | 
| 61784 | 117 | None \<Rightarrow> N1 l | | 
| 118 | Some (b, r') \<Rightarrow> n2 l b r'))" | |
| 119 | ||
| 120 | fun tree :: "'a bro \<Rightarrow> 'a bro" where | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 121 | "tree (N1 t) = t" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 122 | "tree t = t" | 
| 61784 | 123 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 124 | definition delete :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 125 | "delete a t = tree (del a t)" | 
| 61784 | 126 | |
| 127 | end | |
| 128 | ||
| 129 | subsection \<open>Invariants\<close> | |
| 130 | ||
| 131 | fun B :: "nat \<Rightarrow> 'a bro set" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 132 | and U :: "nat \<Rightarrow> 'a bro set" where | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 133 |   "B 0 = {N0}" |
 | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 134 |   "B (Suc h) = { N2 t1 a t2 | t1 a t2. 
 | 
| 61784 | 135 | t1 \<in> B h \<union> U h \<and> t2 \<in> B h \<or> t1 \<in> B h \<and> t2 \<in> B h \<union> U h}" | | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 136 |   "U 0 = {}" |
 | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 137 | "U (Suc h) = N1 ` B h" | 
| 61784 | 138 | |
| 139 | abbreviation "T h \<equiv> B h \<union> U h" | |
| 140 | ||
| 141 | fun Bp :: "nat \<Rightarrow> 'a bro set" where | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 142 | "Bp 0 = B 0 \<union> L2 ` UNIV" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 143 |   "Bp (Suc 0) = B (Suc 0) \<union> {N3 N0 a N0 b N0|a b. True}" |
 | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 144 | "Bp (Suc(Suc h)) = B (Suc(Suc h)) \<union> | 
| 61784 | 145 |   {N3 t1 a t2 b t3 | t1 a t2 b t3. t1 \<in> B (Suc h) \<and> t2 \<in> U (Suc h) \<and> t3 \<in> B (Suc h)}"
 | 
| 146 | ||
| 147 | fun Um :: "nat \<Rightarrow> 'a bro set" where | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 148 |   "Um 0 = {}" |
 | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 149 | "Um (Suc h) = N1 ` T h" | 
| 61784 | 150 | |
| 151 | ||
| 152 | subsection "Functional Correctness Proofs" | |
| 153 | ||
| 154 | subsubsection "Proofs for isin" | |
| 155 | ||
| 67929 | 156 | lemma isin_set: | 
| 157 | "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set(inorder t))" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 158 | by(induction h arbitrary: t) (fastforce simp: isin_simps split: if_splits)+ | 
| 61784 | 159 | |
| 160 | subsubsection "Proofs for insertion" | |
| 161 | ||
| 162 | lemma inorder_n1: "inorder(n1 t) = inorder t" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 163 | by(cases t rule: n1.cases) (auto simp: sorted_lems) | 
| 61784 | 164 | |
| 165 | context insert | |
| 166 | begin | |
| 167 | ||
| 168 | lemma inorder_n2: "inorder(n2 l a r) = inorder l @ a # inorder r" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 169 | by(cases "(l,a,r)" rule: n2.cases) (auto simp: sorted_lems) | 
| 61784 | 170 | |
| 171 | lemma inorder_tree: "inorder(tree t) = inorder t" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 172 | by(cases t) auto | 
| 61784 | 173 | |
| 174 | lemma inorder_ins: "t \<in> T h \<Longrightarrow> | |
| 175 | sorted(inorder t) \<Longrightarrow> inorder(ins a t) = ins_list a (inorder t)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 176 | by(induction h arbitrary: t) (auto simp: ins_list_simps inorder_n1 inorder_n2) | 
| 61784 | 177 | |
| 178 | lemma inorder_insert: "t \<in> T h \<Longrightarrow> | |
| 179 | sorted(inorder t) \<Longrightarrow> inorder(insert a t) = ins_list a (inorder t)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 180 | by(simp add: insert_def inorder_ins inorder_tree) | 
| 61784 | 181 | |
| 182 | end | |
| 183 | ||
| 184 | subsubsection \<open>Proofs for deletion\<close> | |
| 185 | ||
| 186 | context delete | |
| 187 | begin | |
| 188 | ||
| 189 | lemma inorder_tree: "inorder(tree t) = inorder t" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 190 | by(cases t) auto | 
| 61784 | 191 | |
| 192 | lemma inorder_n2: "inorder(n2 l a r) = inorder l @ a # inorder r" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 193 | by(cases "(l,a,r)" rule: n2.cases) (auto) | 
| 61784 | 194 | |
| 68020 | 195 | lemma inorder_split_min: | 
| 196 | "t \<in> T h \<Longrightarrow> (split_min t = None \<longleftrightarrow> inorder t = []) \<and> | |
| 197 | (split_min t = Some(a,t') \<longrightarrow> inorder t = a # inorder t')" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 198 | by(induction h arbitrary: t a t') (auto simp: inorder_n2 split: option.splits) | 
| 61784 | 199 | |
| 200 | lemma inorder_del: | |
| 61792 | 201 | "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)" | 
| 73526 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 202 | apply (induction h arbitrary: t) | 
| 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 203 | apply (auto simp: del_list_simps inorder_n2 split: option.splits) | 
| 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 204 | apply (auto simp: del_list_simps inorder_n2 | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 205 | inorder_split_min[OF UnI1] inorder_split_min[OF UnI2] split: option.splits) | 
| 73526 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
68431diff
changeset | 206 | done | 
| 61792 | 207 | |
| 208 | lemma inorder_delete: | |
| 209 | "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 210 | by(simp add: delete_def inorder_del inorder_tree) | 
| 61784 | 211 | |
| 212 | end | |
| 213 | ||
| 214 | ||
| 215 | subsection \<open>Invariant Proofs\<close> | |
| 216 | ||
| 61789 | 217 | subsubsection \<open>Proofs for insertion\<close> | 
| 61784 | 218 | |
| 219 | lemma n1_type: "t \<in> Bp h \<Longrightarrow> n1 t \<in> T (Suc h)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 220 | by(cases h rule: Bp.cases) auto | 
| 61784 | 221 | |
| 222 | context insert | |
| 223 | begin | |
| 224 | ||
| 61809 | 225 | lemma tree_type: "t \<in> Bp h \<Longrightarrow> tree t \<in> B h \<union> B (Suc h)" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 226 | by(cases h rule: Bp.cases) auto | 
| 61784 | 227 | |
| 228 | lemma n2_type: | |
| 229 | "(t1 \<in> Bp h \<and> t2 \<in> T h \<longrightarrow> n2 t1 a t2 \<in> Bp (Suc h)) \<and> | |
| 230 | (t1 \<in> T h \<and> t2 \<in> Bp h \<longrightarrow> n2 t1 a t2 \<in> Bp (Suc h))" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 231 | apply(cases h rule: Bp.cases) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 232 | apply (auto)[2] | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 233 | apply(rule conjI impI | erule conjE exE imageE | simp | erule disjE)+ | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 234 | done | 
| 61784 | 235 | |
| 236 | lemma Bp_if_B: "t \<in> B h \<Longrightarrow> t \<in> Bp h" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 237 | by (cases h rule: Bp.cases) simp_all | 
| 61784 | 238 | |
| 67406 | 239 | text\<open>An automatic proof:\<close> | 
| 61784 | 240 | |
| 241 | lemma | |
| 242 | "(t \<in> B h \<longrightarrow> ins x t \<in> Bp h) \<and> (t \<in> U h \<longrightarrow> ins x t \<in> T h)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 243 | proof (induction h arbitrary: t) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 244 | case 0 | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 245 | then show ?case by simp | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 246 | next | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 247 | case (Suc h) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 248 | then show ?case by (fastforce simp: Bp_if_B n2_type dest: n1_type) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 249 | qed | 
| 61784 | 250 | |
| 67406 | 251 | text\<open>A detailed proof:\<close> | 
| 61784 | 252 | |
| 253 | lemma ins_type: | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 254 | shows "t \<in> B h \<Longrightarrow> ins x t \<in> Bp h" and "t \<in> U h \<Longrightarrow> ins x t \<in> T h" | 
| 61784 | 255 | proof(induction h arbitrary: t) | 
| 256 | case 0 | |
| 257 |   { case 1 thus ?case by simp
 | |
| 258 | next | |
| 259 | case 2 thus ?case by simp } | |
| 260 | next | |
| 261 | case (Suc h) | |
| 262 |   { case 1
 | |
| 263 | then obtain t1 a t2 where [simp]: "t = N2 t1 a t2" and | |
| 264 | t1: "t1 \<in> T h" and t2: "t2 \<in> T h" and t12: "t1 \<in> B h \<or> t2 \<in> B h" | |
| 265 | by auto | |
| 67040 | 266 | have ?case if "x < a" | 
| 267 | proof - | |
| 268 | have "n2 (ins x t1) a t2 \<in> Bp (Suc h)" | |
| 61784 | 269 | proof cases | 
| 270 | assume "t1 \<in> B h" | |
| 271 | with t2 show ?thesis by (simp add: Suc.IH(1) n2_type) | |
| 272 | next | |
| 273 | assume "t1 \<notin> B h" | |
| 274 | hence 1: "t1 \<in> U h" and 2: "t2 \<in> B h" using t1 t12 by auto | |
| 275 | show ?thesis by (metis Suc.IH(2)[OF 1] Bp_if_B[OF 2] n2_type) | |
| 276 | qed | |
| 67406 | 277 | with \<open>x < a\<close> show ?case by simp | 
| 67040 | 278 | qed | 
| 61784 | 279 | moreover | 
| 67040 | 280 | have ?case if "a < x" | 
| 281 | proof - | |
| 282 | have "n2 t1 a (ins x t2) \<in> Bp (Suc h)" | |
| 61784 | 283 | proof cases | 
| 284 | assume "t2 \<in> B h" | |
| 285 | with t1 show ?thesis by (simp add: Suc.IH(1) n2_type) | |
| 286 | next | |
| 287 | assume "t2 \<notin> B h" | |
| 288 | hence 1: "t1 \<in> B h" and 2: "t2 \<in> U h" using t2 t12 by auto | |
| 289 | show ?thesis by (metis Bp_if_B[OF 1] Suc.IH(2)[OF 2] n2_type) | |
| 290 | qed | |
| 67406 | 291 | with \<open>a < x\<close> show ?case by simp | 
| 67040 | 292 | qed | 
| 293 | moreover | |
| 294 | have ?case if "x = a" | |
| 295 | proof - | |
| 61784 | 296 | from 1 have "t \<in> Bp (Suc h)" by(rule Bp_if_B) | 
| 67406 | 297 | thus "?case" using \<open>x = a\<close> by simp | 
| 67040 | 298 | qed | 
| 61784 | 299 | ultimately show ?case by auto | 
| 300 | next | |
| 301 | case 2 thus ?case using Suc(1) n1_type by fastforce } | |
| 302 | qed | |
| 303 | ||
| 304 | lemma insert_type: | |
| 61809 | 305 | "t \<in> B h \<Longrightarrow> insert x t \<in> B h \<union> B (Suc h)" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 306 | unfolding insert_def by (metis ins_type(1) tree_type) | 
| 61784 | 307 | |
| 308 | end | |
| 309 | ||
| 61789 | 310 | subsubsection "Proofs for deletion" | 
| 61784 | 311 | |
| 312 | lemma B_simps[simp]: | |
| 313 | "N1 t \<in> B h = False" | |
| 314 | "L2 y \<in> B h = False" | |
| 315 | "(N3 t1 a1 t2 a2 t3) \<in> B h = False" | |
| 316 | "N0 \<in> B h \<longleftrightarrow> h = 0" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 317 | by (cases h, auto)+ | 
| 61784 | 318 | |
| 319 | context delete | |
| 320 | begin | |
| 321 | ||
| 322 | lemma n2_type1: | |
| 323 | "\<lbrakk>t1 \<in> Um h; t2 \<in> B h\<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 324 | apply(cases h rule: Bp.cases) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 325 | apply auto[2] | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 326 | apply(erule exE bexE conjE imageE | simp | erule disjE)+ | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 327 | done | 
| 61784 | 328 | |
| 329 | lemma n2_type2: | |
| 330 | "\<lbrakk>t1 \<in> B h ; t2 \<in> Um h \<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 331 | apply(cases h rule: Bp.cases) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 332 | using Um.simps(1) apply blast | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 333 | apply force | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 334 | apply(erule exE bexE conjE imageE | simp | erule disjE)+ | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 335 | done | 
| 61784 | 336 | |
| 337 | lemma n2_type3: | |
| 338 | "\<lbrakk>t1 \<in> T h ; t2 \<in> T h \<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)" | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 339 | apply(cases h rule: Bp.cases) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 340 | apply auto[2] | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 341 | apply(erule exE bexE conjE imageE | simp | erule disjE)+ | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 342 | done | 
| 61784 | 343 | |
| 68020 | 344 | lemma split_minNoneN0: "\<lbrakk>t \<in> B h; split_min t = None\<rbrakk> \<Longrightarrow> t = N0" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 345 | by (cases t) (auto split: option.splits) | 
| 61784 | 346 | |
| 68020 | 347 | lemma split_minNoneN1 : "\<lbrakk>t \<in> U h; split_min t = None\<rbrakk> \<Longrightarrow> t = N1 N0" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 348 | by (cases h) (auto simp: split_minNoneN0 split: option.splits) | 
| 61784 | 349 | |
| 68020 | 350 | lemma split_min_type: | 
| 351 | "t \<in> B h \<Longrightarrow> split_min t = Some (a, t') \<Longrightarrow> t' \<in> T h" | |
| 352 | "t \<in> U h \<Longrightarrow> split_min t = Some (a, t') \<Longrightarrow> t' \<in> Um h" | |
| 61784 | 353 | proof (induction h arbitrary: t a t') | 
| 354 | case (Suc h) | |
| 355 |   { case 1
 | |
| 356 | then obtain t1 a t2 where [simp]: "t = N2 t1 a t2" and | |
| 357 | t12: "t1 \<in> T h" "t2 \<in> T h" "t1 \<in> B h \<or> t2 \<in> B h" | |
| 358 | by auto | |
| 359 | show ?case | |
| 68020 | 360 | proof (cases "split_min t1") | 
| 61784 | 361 | case None | 
| 362 | show ?thesis | |
| 363 | proof cases | |
| 364 | assume "t1 \<in> B h" | |
| 68020 | 365 | with split_minNoneN0[OF this None] 1 show ?thesis by(auto) | 
| 61784 | 366 | next | 
| 367 | assume "t1 \<notin> B h" | |
| 368 | thus ?thesis using 1 None by (auto) | |
| 369 | qed | |
| 370 | next | |
| 371 | case [simp]: (Some bt') | |
| 372 | obtain b t1' where [simp]: "bt' = (b,t1')" by fastforce | |
| 373 | show ?thesis | |
| 374 | proof cases | |
| 375 | assume "t1 \<in> B h" | |
| 376 | from Suc.IH(1)[OF this] 1 have "t1' \<in> T h" by simp | |
| 377 | from n2_type3[OF this t12(2)] 1 show ?thesis by auto | |
| 378 | next | |
| 379 | assume "t1 \<notin> B h" | |
| 380 | hence t1: "t1 \<in> U h" and t2: "t2 \<in> B h" using t12 by auto | |
| 381 | from Suc.IH(2)[OF t1] have "t1' \<in> Um h" by simp | |
| 382 | from n2_type1[OF this t2] 1 show ?thesis by auto | |
| 383 | qed | |
| 384 | qed | |
| 385 | } | |
| 386 |   { case 2
 | |
| 387 | then obtain t1 where [simp]: "t = N1 t1" and t1: "t1 \<in> B h" by auto | |
| 388 | show ?case | |
| 68020 | 389 | proof (cases "split_min t1") | 
| 61784 | 390 | case None | 
| 68020 | 391 | with split_minNoneN0[OF t1 None] 2 show ?thesis by(auto) | 
| 61784 | 392 | next | 
| 393 | case [simp]: (Some bt') | |
| 394 | obtain b t1' where [simp]: "bt' = (b,t1')" by fastforce | |
| 395 | from Suc.IH(1)[OF t1] have "t1' \<in> T h" by simp | |
| 396 | thus ?thesis using 2 by auto | |
| 397 | qed | |
| 398 | } | |
| 399 | qed auto | |
| 400 | ||
| 401 | lemma del_type: | |
| 402 | "t \<in> B h \<Longrightarrow> del x t \<in> T h" | |
| 403 | "t \<in> U h \<Longrightarrow> del x t \<in> Um h" | |
| 404 | proof (induction h arbitrary: x t) | |
| 405 | case (Suc h) | |
| 406 |   { case 1
 | |
| 407 | then obtain l a r where [simp]: "t = N2 l a r" and | |
| 408 | lr: "l \<in> T h" "r \<in> T h" "l \<in> B h \<or> r \<in> B h" by auto | |
| 67040 | 409 | have ?case if "x < a" | 
| 410 | proof cases | |
| 411 | assume "l \<in> B h" | |
| 412 | from n2_type3[OF Suc.IH(1)[OF this] lr(2)] | |
| 67406 | 413 | show ?thesis using \<open>x<a\<close> by(simp) | 
| 67040 | 414 | next | 
| 415 | assume "l \<notin> B h" | |
| 416 | hence "l \<in> U h" "r \<in> B h" using lr by auto | |
| 417 | from n2_type1[OF Suc.IH(2)[OF this(1)] this(2)] | |
| 67406 | 418 | show ?thesis using \<open>x<a\<close> by(simp) | 
| 67040 | 419 | qed | 
| 420 | moreover | |
| 421 | have ?case if "x > a" | |
| 422 | proof cases | |
| 423 | assume "r \<in> B h" | |
| 424 | from n2_type3[OF lr(1) Suc.IH(1)[OF this]] | |
| 67406 | 425 | show ?thesis using \<open>x>a\<close> by(simp) | 
| 67040 | 426 | next | 
| 427 | assume "r \<notin> B h" | |
| 428 | hence "l \<in> B h" "r \<in> U h" using lr by auto | |
| 429 | from n2_type2[OF this(1) Suc.IH(2)[OF this(2)]] | |
| 67406 | 430 | show ?thesis using \<open>x>a\<close> by(simp) | 
| 67040 | 431 | qed | 
| 432 | moreover | |
| 433 | have ?case if [simp]: "x=a" | |
| 68020 | 434 | proof (cases "split_min r") | 
| 67040 | 435 | case None | 
| 436 | show ?thesis | |
| 61784 | 437 | proof cases | 
| 438 | assume "r \<in> B h" | |
| 68020 | 439 | with split_minNoneN0[OF this None] lr show ?thesis by(simp) | 
| 61784 | 440 | next | 
| 441 | assume "r \<notin> B h" | |
| 67040 | 442 | hence "r \<in> U h" using lr by auto | 
| 68020 | 443 | with split_minNoneN1[OF this None] lr(3) show ?thesis by (simp) | 
| 61784 | 444 | qed | 
| 67040 | 445 | next | 
| 446 | case [simp]: (Some br') | |
| 447 | obtain b r' where [simp]: "br' = (b,r')" by fastforce | |
| 448 | show ?thesis | |
| 449 | proof cases | |
| 450 | assume "r \<in> B h" | |
| 68020 | 451 | from split_min_type(1)[OF this] n2_type3[OF lr(1)] | 
| 67040 | 452 | show ?thesis by simp | 
| 61784 | 453 | next | 
| 67040 | 454 | assume "r \<notin> B h" | 
| 455 | hence "l \<in> B h" and "r \<in> U h" using lr by auto | |
| 68020 | 456 | from split_min_type(2)[OF this(2)] n2_type2[OF this(1)] | 
| 67040 | 457 | show ?thesis by simp | 
| 61784 | 458 | qed | 
| 67040 | 459 | qed | 
| 460 | ultimately show ?case by auto | |
| 61784 | 461 | } | 
| 462 |   { case 2 with Suc.IH(1) show ?case by auto }
 | |
| 463 | qed auto | |
| 464 | ||
| 67613 | 465 | lemma tree_type: "t \<in> T (h+1) \<Longrightarrow> tree t \<in> B (h+1) \<union> B h" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 466 | by(auto) | 
| 61784 | 467 | |
| 61809 | 468 | lemma delete_type: "t \<in> B h \<Longrightarrow> delete x t \<in> B h \<union> B(h-1)" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 469 | unfolding delete_def | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 470 | by (cases h) (simp, metis del_type(1) tree_type Suc_eq_plus1 diff_Suc_1) | 
| 61784 | 471 | |
| 472 | end | |
| 473 | ||
| 61789 | 474 | |
| 61784 | 475 | subsection "Overall correctness" | 
| 476 | ||
| 477 | interpretation Set_by_Ordered | |
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 478 | where empty = empty and isin = isin and insert = insert.insert | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 479 | and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> B h" | 
| 61784 | 480 | proof (standard, goal_cases) | 
| 481 | case 2 thus ?case by(auto intro!: isin_set) | |
| 482 | next | |
| 483 | case 3 thus ?case by(auto intro!: insert.inorder_insert) | |
| 484 | next | |
| 61792 | 485 | case 4 thus ?case by(auto intro!: delete.inorder_delete) | 
| 61784 | 486 | next | 
| 487 | case 6 thus ?case using insert.insert_type by blast | |
| 488 | next | |
| 489 | case 7 thus ?case using delete.delete_type by blast | |
| 68431 | 490 | qed (auto simp: empty_def) | 
| 61784 | 491 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 492 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 493 | subsection \<open>Height-Size Relation\<close> | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 494 | |
| 76063 | 495 | text \<open>By Daniel Stüwe\<close> | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 496 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 497 | fun fib_tree :: "nat \<Rightarrow> unit bro" where | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 498 | "fib_tree 0 = N0" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 499 | | "fib_tree (Suc 0) = N2 N0 () N0" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 500 | | "fib_tree (Suc(Suc h)) = N2 (fib_tree (h+1)) () (N1 (fib_tree h))" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 501 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 502 | fun fib' :: "nat \<Rightarrow> nat" where | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 503 | "fib' 0 = 0" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 504 | | "fib' (Suc 0) = 1" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 505 | | "fib' (Suc(Suc h)) = 1 + fib' (Suc h) + fib' h" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 506 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 507 | fun size :: "'a bro \<Rightarrow> nat" where | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 508 | "size N0 = 0" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 509 | | "size (N1 t) = size t" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 510 | | "size (N2 t1 _ t2) = 1 + size t1 + size t2" | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 511 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 512 | lemma fib_tree_B: "fib_tree h \<in> B h" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 513 | by (induction h rule: fib_tree.induct) auto | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 514 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 515 | declare [[names_short]] | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 516 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 517 | lemma size_fib': "size (fib_tree h) = fib' h" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 518 | by (induction h rule: fib_tree.induct) auto | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 519 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 520 | lemma fibfib: "fib' h + 1 = fib (Suc(Suc h))" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 521 | by (induction h rule: fib_tree.induct) auto | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 522 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 523 | lemma B_N2_cases[consumes 1]: | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 524 | assumes "N2 t1 a t2 \<in> B (Suc n)" | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 525 | obtains | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 526 | (BB) "t1 \<in> B n" and "t2 \<in> B n" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 527 | (UB) "t1 \<in> U n" and "t2 \<in> B n" | | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 528 | (BU) "t1 \<in> B n" and "t2 \<in> U n" | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 529 | using assms by auto | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 530 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 531 | lemma size_bounded: "t \<in> B h \<Longrightarrow> size t \<ge> size (fib_tree h)" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 532 | unfolding size_fib' proof (induction h arbitrary: t rule: fib'.induct) | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 533 | case (3 h t') | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 534 | note main = 3 | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 535 | then obtain t1 a t2 where t': "t' = N2 t1 a t2" by auto | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 536 | with main have "N2 t1 a t2 \<in> B (Suc (Suc h))" by auto | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 537 | thus ?case proof (cases rule: B_N2_cases) | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 538 | case BB | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 539 | then obtain x y z where t2: "t2 = N2 x y z \<or> t2 = N2 z y x" "x \<in> B h" by auto | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 540 | show ?thesis unfolding t' using main(1)[OF BB(1)] main(2)[OF t2(2)] t2(1) by auto | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 541 | next | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 542 | case UB | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 543 | then obtain t11 where t1: "t1 = N1 t11" "t11 \<in> B h" by auto | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 544 | show ?thesis unfolding t' t1(1) using main(2)[OF t1(2)] main(1)[OF UB(2)] by simp | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 545 | next | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 546 | case BU | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 547 | then obtain t22 where t2: "t2 = N1 t22" "t22 \<in> B h" by auto | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 548 | show ?thesis unfolding t' t2(1) using main(2)[OF t2(2)] main(1)[OF BU(1)] by simp | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 549 | qed | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 550 | qed auto | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 551 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 552 | theorem "t \<in> B h \<Longrightarrow> fib (h + 2) \<le> size t + 1" | 
| 78653 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 553 | using size_bounded | 
| 
7ed1759fe1bd
tidying up old apply-style proofs
 paulson <lp15@cam.ac.uk> parents: 
76063diff
changeset | 554 | by (simp add: size_fib' fibfib[symmetric] del: fib.simps) | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 555 | |
| 61784 | 556 | end |