src/HOL/Data_Structures/Time_Funs.thy
changeset 81357 21a493abde0f
parent 81355 8070e4578ece
--- a/src/HOL/Data_Structures/Time_Funs.thy	Wed Nov 06 16:27:06 2024 +0100
+++ b/src/HOL/Data_Structures/Time_Funs.thy	Wed Nov 06 18:10:39 2024 +0100
@@ -29,7 +29,7 @@
 abbreviation T_length :: "'a list \<Rightarrow> nat" where
 "T_length \<equiv> T_size"
 
-lemma T_length_eq: "T_length xs = length xs + 1"
+lemma T_length: "T_length xs = length xs + 1"
   by (induction xs) auto
 
 lemmas [simp del] = T_size_list.simps
@@ -41,7 +41,7 @@
   "T_map T_f (x # xs) = T_f x + T_map T_f xs + 1"
 by (simp_all add: T_map_def)
 
-lemma T_map_eq: "T_map T_f xs = (\<Sum>x\<leftarrow>xs. T_f x) + length xs + 1"
+lemma T_map: "T_map T_f xs = (\<Sum>x\<leftarrow>xs. T_f x) + length xs + 1"
   by (induction xs) auto
 
 lemmas [simp del] = T_map_simps
@@ -53,12 +53,12 @@
   "T_filter T_P (x # xs) = T_P x + T_filter T_P xs + 1"
 by (simp_all add: T_filter_def)
 
-lemma T_filter_eq: "T_filter T_P xs = (\<Sum>x\<leftarrow>xs. T_P x) + length xs + 1"
+lemma T_filter: "T_filter T_P xs = (\<Sum>x\<leftarrow>xs. T_P x) + length xs + 1"
 by (induction xs) (auto simp: T_filter_simps)
 
 time_fun nth
 
-lemma T_nth_eq: "n < length xs \<Longrightarrow> T_nth xs n = n + 1"
+lemma T_nth: "n < length xs \<Longrightarrow> T_nth xs n = n + 1"
   by (induction xs n rule: T_nth.induct) (auto split: nat.splits)
 
 lemmas [simp del] = T_nth.simps
@@ -66,10 +66,10 @@
 time_fun take
 time_fun drop
 
-lemma T_take_eq: "T_take n xs = min n (length xs) + 1"
+lemma T_take: "T_take n xs = min n (length xs) + 1"
   by (induction xs arbitrary: n) (auto split: nat.splits)
 
-lemma T_drop_eq: "T_drop n xs = min n (length xs) + 1"
+lemma T_drop: "T_drop n xs = min n (length xs) + 1"
   by (induction xs arbitrary: n) (auto split: nat.splits)
 
 time_fun rev