# HG changeset patch # User nipkow # Date 1705680877 -3600 # Node ID 8a25110626095d00e25490f1bc3f9d5b8fdce8bf # Parent c7536609bb9b8945e94e9fc0e497dfd33dfa66a4 more uses of define_time_fun diff -r c7536609bb9b -r 8a2511062609 src/HOL/Data_Structures/Queue_2Lists.thy --- a/src/HOL/Data_Structures/Queue_2Lists.thy Thu Jan 18 14:30:27 2024 +0100 +++ b/src/HOL/Data_Structures/Queue_2Lists.thy Fri Jan 19 17:14:37 2024 +0100 @@ -59,14 +59,13 @@ text \Running times:\ -fun T_norm :: "'a queue \ nat" where -"T_norm (fs,rs) = (if fs = [] then T_itrev rs [] else 0)" +define_time_fun norm +define_time_fun enq +define_time_fun tl +define_time_fun deq -fun T_enq :: "'a \ 'a queue \ nat" where -"T_enq a (fs,rs) = T_norm(fs, a # rs)" - -fun T_deq :: "'a queue \ nat" where -"T_deq (fs,rs) = (if fs = [] then 0 else T_norm(tl fs,rs))" +lemma T_tl_0: "T_tl xs = 0" +by(cases xs)auto text \Amortized running times:\ @@ -77,6 +76,6 @@ by(auto simp: T_itrev) lemma a_deq: "T_deq (fs,rs) + \(deq (fs,rs)) - \(fs,rs) \ 1" -by(auto simp: T_itrev) +by(auto simp: T_itrev T_tl_0) end diff -r c7536609bb9b -r 8a2511062609 src/HOL/Data_Structures/Reverse.thy --- a/src/HOL/Data_Structures/Reverse.thy Thu Jan 18 14:30:27 2024 +0100 +++ b/src/HOL/Data_Structures/Reverse.thy Fri Jan 19 17:14:37 2024 +0100 @@ -1,14 +1,9 @@ theory Reverse -imports Main +imports Define_Time_Function begin -fun T_append :: "'a list \ 'a list \ nat" where -"T_append [] ys = 1" | -"T_append (x#xs) ys = T_append xs ys + 1" - -fun T_rev :: "'a list \ nat" where -"T_rev [] = 1" | -"T_rev (x#xs) = T_rev xs + T_append (rev xs) [x] + 1" +define_time_fun append +define_time_fun rev lemma T_append: "T_append xs ys = length xs + 1" by(induction xs) auto @@ -26,9 +21,7 @@ lemma itrev_Nil: "itrev xs [] = rev xs" by(simp add: itrev) -fun T_itrev :: "'a list \ 'a list \ nat" where -"T_itrev [] ys = 1" | -"T_itrev (x#xs) ys = T_itrev xs (x # ys) + 1" +define_time_fun itrev lemma T_itrev: "T_itrev xs ys = length xs + 1" by(induction xs arbitrary: ys) auto diff -r c7536609bb9b -r 8a2511062609 src/HOL/Data_Structures/Time_Funs.thy --- a/src/HOL/Data_Structures/Time_Funs.thy Thu Jan 18 14:30:27 2024 +0100 +++ b/src/HOL/Data_Structures/Time_Funs.thy Fri Jan 19 17:14:37 2024 +0100 @@ -4,9 +4,11 @@ *) section \Time functions for various standard library operations\ theory Time_Funs - imports Main + imports Define_Time_Function begin +text \Automatic definition of \T_length\ is cumbersome because of the type class for \size\.\ + fun T_length :: "'a list \ nat" where "T_length [] = 1" | "T_length (x # xs) = T_length xs + 1" @@ -46,20 +48,13 @@ lemmas [simp del] = T_nth.simps - -fun T_take :: "nat \ 'a list \ nat" where - "T_take n [] = 1" -| "T_take n (x # xs) = (case n of 0 \ 1 | Suc n' \ T_take n' xs + 1)" +define_time_fun take +define_time_fun drop lemma T_take_eq: "T_take n xs = min n (length xs) + 1" by (induction xs arbitrary: n) (auto split: nat.splits) -fun T_drop :: "nat \ 'a list \ nat" where - "T_drop n [] = 1" -| "T_drop n (x # xs) = (case n of 0 \ 1 | Suc n' \ T_drop n' xs + 1)" - lemma T_drop_eq: "T_drop n xs = min n (length xs) + 1" by (induction xs arbitrary: n) (auto split: nat.splits) - end