| 72642 |      1 | theory Reverse
 | 
|  |      2 | imports Main
 | 
|  |      3 | begin
 | 
|  |      4 | 
 | 
|  |      5 | fun T_append :: "'a list \<Rightarrow> 'a list \<Rightarrow> nat" where
 | 
|  |      6 | "T_append [] ys = 1" |
 | 
|  |      7 | "T_append (x#xs) ys = T_append xs ys + 1"
 | 
|  |      8 | 
 | 
|  |      9 | fun T_rev :: "'a list \<Rightarrow> nat" where
 | 
|  |     10 | "T_rev [] = 1" |
 | 
|  |     11 | "T_rev (x#xs) = T_rev xs + T_append (rev xs) [x] + 1"
 | 
|  |     12 | 
 | 
|  |     13 | lemma T_append: "T_append xs ys = length xs + 1"
 | 
|  |     14 | by(induction xs) auto
 | 
|  |     15 | 
 | 
|  |     16 | lemma T_rev: "T_rev xs \<le> (length xs + 1)^2"
 | 
|  |     17 | by(induction xs) (auto simp: T_append power2_eq_square)
 | 
|  |     18 | 
 | 
|  |     19 | fun itrev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
 | 
|  |     20 | "itrev [] ys = ys" |
 | 
|  |     21 | "itrev (x#xs) ys = itrev xs (x # ys)"
 | 
|  |     22 | 
 | 
|  |     23 | lemma itrev: "itrev xs ys = rev xs @ ys"
 | 
|  |     24 | by(induction xs arbitrary: ys) auto
 | 
|  |     25 | 
 | 
|  |     26 | lemma itrev_Nil: "itrev xs [] = rev xs"
 | 
|  |     27 | by(simp add: itrev)
 | 
|  |     28 | 
 | 
|  |     29 | fun T_itrev :: "'a list \<Rightarrow> 'a list \<Rightarrow> nat" where
 | 
|  |     30 | "T_itrev [] ys = 1" |
 | 
|  |     31 | "T_itrev (x#xs) ys = T_itrev xs (x # ys) + 1"
 | 
|  |     32 | 
 | 
|  |     33 | lemma T_itrev: "T_itrev xs ys = length xs + 1"
 | 
|  |     34 | by(induction xs arbitrary: ys) auto
 | 
|  |     35 | 
 | 
|  |     36 | end |