src/HOL/Data_Structures/Time_Funs.thy
changeset 73108 981a383610df
child 79495 8a2511062609
equal deleted inserted replaced
73107:f062d19c4b44 73108:981a383610df
       
     1 (*
       
     2   File:    Data_Structures/Time_Functions.thy
       
     3   Author:  Manuel Eberl, TU München
       
     4 *)
       
     5 section \<open>Time functions for various standard library operations\<close>
       
     6 theory Time_Funs
       
     7   imports Main
       
     8 begin
       
     9 
       
    10 fun T_length :: "'a list \<Rightarrow> nat" where
       
    11   "T_length [] = 1"
       
    12 | "T_length (x # xs) = T_length xs + 1"
       
    13 
       
    14 lemma T_length_eq: "T_length xs = length xs + 1"
       
    15   by (induction xs) auto
       
    16 
       
    17 lemmas [simp del] = T_length.simps
       
    18 
       
    19 
       
    20 fun T_map  :: "('a \<Rightarrow> nat) \<Rightarrow> 'a list \<Rightarrow> nat" where
       
    21   "T_map T_f [] = 1"
       
    22 | "T_map T_f (x # xs) = T_f x + T_map T_f xs + 1"
       
    23 
       
    24 lemma T_map_eq: "T_map T_f xs = (\<Sum>x\<leftarrow>xs. T_f x) + length xs + 1"
       
    25   by (induction xs) auto
       
    26 
       
    27 lemmas [simp del] = T_map.simps
       
    28 
       
    29 
       
    30 fun T_filter  :: "('a \<Rightarrow> nat) \<Rightarrow> 'a list \<Rightarrow> nat" where
       
    31   "T_filter T_p [] = 1"
       
    32 | "T_filter T_p (x # xs) = T_p x + T_filter T_p xs + 1"
       
    33 
       
    34 lemma T_filter_eq: "T_filter T_p xs = (\<Sum>x\<leftarrow>xs. T_p x) + length xs + 1"
       
    35   by (induction xs) auto
       
    36 
       
    37 lemmas [simp del] = T_filter.simps
       
    38 
       
    39 
       
    40 fun T_nth :: "'a list \<Rightarrow> nat \<Rightarrow> nat" where
       
    41   "T_nth [] n = 1"
       
    42 | "T_nth (x # xs) n = (case n of 0 \<Rightarrow> 1 | Suc n' \<Rightarrow> T_nth xs n' + 1)"
       
    43 
       
    44 lemma T_nth_eq: "T_nth xs n = min n (length xs) + 1"
       
    45   by (induction xs n rule: T_nth.induct) (auto split: nat.splits)
       
    46 
       
    47 lemmas [simp del] = T_nth.simps
       
    48 
       
    49 
       
    50 fun T_take :: "nat \<Rightarrow> 'a list \<Rightarrow> nat" where
       
    51   "T_take n [] = 1"
       
    52 | "T_take n (x # xs) = (case n of 0 \<Rightarrow> 1 | Suc n' \<Rightarrow> T_take n' xs + 1)"
       
    53 
       
    54 lemma T_take_eq: "T_take n xs = min n (length xs) + 1"
       
    55   by (induction xs arbitrary: n) (auto split: nat.splits)
       
    56 
       
    57 fun T_drop :: "nat \<Rightarrow> 'a list \<Rightarrow> nat" where
       
    58   "T_drop n [] = 1"
       
    59 | "T_drop n (x # xs) = (case n of 0 \<Rightarrow> 1 | Suc n' \<Rightarrow> T_drop n' xs + 1)"
       
    60 
       
    61 lemma T_drop_eq: "T_drop n xs = min n (length xs) + 1"
       
    62   by (induction xs arbitrary: n) (auto split: nat.splits)
       
    63   
       
    64   
       
    65 end