more uses of define_time_fun
authornipkow
Fri, 19 Jan 2024 17:14:37 +0100
changeset 79495 8a2511062609
parent 79494 c7536609bb9b
child 79496 4d82b743c5e1
more uses of define_time_fun
src/HOL/Data_Structures/Queue_2Lists.thy
src/HOL/Data_Structures/Reverse.thy
src/HOL/Data_Structures/Time_Funs.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 \<open>Running times:\<close>
 
-fun T_norm :: "'a queue \<Rightarrow> 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 \<Rightarrow> 'a queue \<Rightarrow> nat" where
-"T_enq a (fs,rs) = T_norm(fs, a # rs)"
-
-fun T_deq :: "'a queue \<Rightarrow> 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 \<open>Amortized running times:\<close>
 
@@ -77,6 +76,6 @@
 by(auto simp: T_itrev)
 
 lemma a_deq: "T_deq (fs,rs) + \<Phi>(deq (fs,rs)) - \<Phi>(fs,rs) \<le> 1"
-by(auto simp: T_itrev)
+by(auto simp: T_itrev T_tl_0)
 
 end
--- 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 \<Rightarrow> 'a list \<Rightarrow> nat" where
-"T_append [] ys = 1" |
-"T_append (x#xs) ys = T_append xs ys + 1"
-
-fun T_rev :: "'a list \<Rightarrow> 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 \<Rightarrow> 'a list \<Rightarrow> 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
--- 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 \<open>Time functions for various standard library operations\<close>
 theory Time_Funs
-  imports Main
+  imports Define_Time_Function
 begin
 
+text \<open>Automatic definition of \<open>T_length\<close> is cumbersome because of the type class for \<open>size\<close>.\<close>
+
 fun T_length :: "'a list \<Rightarrow> 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 \<Rightarrow> 'a list \<Rightarrow> nat" where
-  "T_take n [] = 1"
-| "T_take n (x # xs) = (case n of 0 \<Rightarrow> 1 | Suc n' \<Rightarrow> 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 \<Rightarrow> 'a list \<Rightarrow> nat" where
-  "T_drop n [] = 1"
-| "T_drop n (x # xs) = (case n of 0 \<Rightarrow> 1 | Suc n' \<Rightarrow> 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