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