| author | wenzelm | 
| Sun, 25 Aug 2024 23:19:33 +0200 | |
| changeset 80772 | 39641d8bd422 | 
| parent 80734 | 7054a1bc8347 | 
| child 81147 | 503e5280ba72 | 
| permissions | -rw-r--r-- | 
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 1 | (* | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 2 | File: Data_Structures/Time_Functions.thy | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 3 | Author: Manuel Eberl, TU München | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 4 | *) | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 5 | section \<open>Time functions for various standard library operations\<close> | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 6 | theory Time_Funs | 
| 79495 | 7 | imports Define_Time_Function | 
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 8 | begin | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 9 | |
| 80036 | 10 | time_fun "(@)" | 
| 11 | ||
| 12 | lemma T_append: "T_append xs ys = length xs + 1" | |
| 13 | by(induction xs) auto | |
| 14 | ||
| 80734 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 15 | class T_size = | 
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 16 | fixes T_size :: "'a \<Rightarrow> nat" | 
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 17 | |
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 18 | instantiation list :: (_) T_size | 
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 19 | begin | 
| 79495 | 20 | |
| 80734 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 21 | time_fun length | 
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 22 | |
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 23 | instance .. | 
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 24 | |
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 25 | end | 
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 26 | |
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 27 | abbreviation T_length :: "'a list \<Rightarrow> nat" where | 
| 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 28 | "T_length \<equiv> T_size" | 
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 29 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 30 | lemma T_length_eq: "T_length xs = length xs + 1" | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 31 | by (induction xs) auto | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 32 | |
| 80734 
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
 nipkow parents: 
80624diff
changeset | 33 | lemmas [simp del] = T_size_list.simps | 
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 34 | |
| 80624 
9f8034d29365
time_function T_map can now be generated automatically.
 nipkow parents: 
80107diff
changeset | 35 | time_fun map | 
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 36 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 37 | lemma T_map_eq: "T_map T_f xs = (\<Sum>x\<leftarrow>xs. T_f x) + length xs + 1" | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 38 | by (induction xs) auto | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 39 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 40 | lemmas [simp del] = T_map.simps | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 41 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 42 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 43 | fun T_filter  :: "('a \<Rightarrow> nat) \<Rightarrow> 'a list \<Rightarrow> nat" where
 | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 44 | "T_filter T_p [] = 1" | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 45 | | "T_filter T_p (x # xs) = T_p x + T_filter T_p xs + 1" | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 46 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 47 | lemma T_filter_eq: "T_filter T_p xs = (\<Sum>x\<leftarrow>xs. T_p x) + length xs + 1" | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 48 | by (induction xs) auto | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 49 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 50 | lemmas [simp del] = T_filter.simps | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 51 | |
| 80107 
247751d25102
canonical time function for List.nth
 Manuel Eberl <eberlm@in.tum.de> parents: 
80036diff
changeset | 52 | time_fun nth | 
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 53 | |
| 80107 
247751d25102
canonical time function for List.nth
 Manuel Eberl <eberlm@in.tum.de> parents: 
80036diff
changeset | 54 | lemma T_nth_eq: "n < length xs \<Longrightarrow> T_nth xs n = n + 1" | 
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 55 | by (induction xs n rule: T_nth.induct) (auto split: nat.splits) | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 56 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 57 | lemmas [simp del] = T_nth.simps | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 58 | |
| 79969 | 59 | time_fun take | 
| 60 | time_fun drop | |
| 73108 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 61 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 62 | lemma T_take_eq: "T_take n xs = min n (length xs) + 1" | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 63 | by (induction xs arbitrary: n) (auto split: nat.splits) | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 64 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 65 | lemma T_drop_eq: "T_drop n xs = min n (length xs) + 1" | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 66 | by (induction xs arbitrary: n) (auto split: nat.splits) | 
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 67 | |
| 
981a383610df
HOL-Data_Structures: added Selection and time functions for list functions
 Manuel Eberl <eberlm@in.tum.de> parents: diff
changeset | 68 | end |