src/HOL/Data_Structures/Time_Funs.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81357 21a493abde0f
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
*)
81355
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
     5
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
     6
section \<open>Time functions for various standard library operations. Also defines \<open>itrev\<close>.\<close>
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
     7
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     8
theory Time_Funs
79495
8a2511062609 more uses of define_time_fun
nipkow
parents: 73108
diff changeset
     9
  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
    10
begin
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    11
80036
a594d22e69d6 updated time functions for Array_Braun
nipkow
parents: 79969
diff changeset
    12
time_fun "(@)"
a594d22e69d6 updated time functions for Array_Braun
nipkow
parents: 79969
diff changeset
    13
a594d22e69d6 updated time functions for Array_Braun
nipkow
parents: 79969
diff changeset
    14
lemma T_append: "T_append xs ys = length xs + 1"
a594d22e69d6 updated time functions for Array_Braun
nipkow
parents: 79969
diff changeset
    15
by(induction xs) auto
a594d22e69d6 updated time functions for Array_Braun
nipkow
parents: 79969
diff changeset
    16
80734
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    17
class T_size =
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    18
  fixes T_size :: "'a \<Rightarrow> nat"
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    19
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    20
instantiation list :: (_) T_size
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    21
begin
79495
8a2511062609 more uses of define_time_fun
nipkow
parents: 73108
diff changeset
    22
80734
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    23
time_fun length
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    24
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    25
instance ..
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    26
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    27
end
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    28
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    29
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: 80624
diff changeset
    30
"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
    31
81357
21a493abde0f uniform name T_f for closed-form lemmas for function T_f
nipkow
parents: 81355
diff changeset
    32
lemma T_length: "T_length xs = length xs + 1"
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    33
  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
    34
80734
7054a1bc8347 new version of time_fun that works for classes; define T_length automatically now
nipkow
parents: 80624
diff changeset
    35
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
    36
80624
9f8034d29365 time_function T_map can now be generated automatically.
nipkow
parents: 80107
diff changeset
    37
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
    38
81147
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    39
lemma T_map_simps [simp,code]:
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    40
  "T_map T_f [] = 1"
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    41
  "T_map T_f (x # xs) = T_f x + T_map T_f xs + 1"
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    42
by (simp_all add: T_map_def)
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    43
81357
21a493abde0f uniform name T_f for closed-form lemmas for function T_f
nipkow
parents: 81355
diff changeset
    44
lemma T_map: "T_map T_f xs = (\<Sum>x\<leftarrow>xs. T_f x) + length xs + 1"
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    45
  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
    46
81147
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    47
lemmas [simp del] = T_map_simps
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    48
81147
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    49
time_fun filter
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    50
81147
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    51
lemma T_filter_simps [code]:
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    52
  "T_filter T_P [] = 1"
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    53
  "T_filter T_P (x # xs) = T_P x + T_filter T_P xs + 1"
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    54
by (simp_all add: T_filter_def)
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    55
81357
21a493abde0f uniform name T_f for closed-form lemmas for function T_f
nipkow
parents: 81355
diff changeset
    56
lemma T_filter: "T_filter T_P xs = (\<Sum>x\<leftarrow>xs. T_P x) + length xs + 1"
81147
503e5280ba72 new HO time functions
nipkow
parents: 80734
diff changeset
    57
by (induction xs) (auto simp: T_filter_simps)
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    58
80107
247751d25102 canonical time function for List.nth
Manuel Eberl <eberlm@in.tum.de>
parents: 80036
diff changeset
    59
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
    60
81357
21a493abde0f uniform name T_f for closed-form lemmas for function T_f
nipkow
parents: 81355
diff changeset
    61
lemma T_nth: "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
    62
  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
    63
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    64
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
    65
79969
4aeb25ba90f3 more uniform command names
nipkow
parents: 79495
diff changeset
    66
time_fun take
4aeb25ba90f3 more uniform command names
nipkow
parents: 79495
diff changeset
    67
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
    68
81357
21a493abde0f uniform name T_f for closed-form lemmas for function T_f
nipkow
parents: 81355
diff changeset
    69
lemma T_take: "T_take n xs = min n (length xs) + 1"
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    70
  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
    71
81357
21a493abde0f uniform name T_f for closed-form lemmas for function T_f
nipkow
parents: 81355
diff changeset
    72
lemma T_drop: "T_drop n xs = min n (length xs) + 1"
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    73
  by (induction xs arbitrary: n) (auto split: nat.splits)
81355
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    74
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    75
time_fun rev
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    76
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    77
lemma T_rev: "T_rev xs \<le> (length xs + 1)^2"
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    78
by(induction xs) (auto simp: T_append power2_eq_square)
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    79
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    80
fun itrev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    81
"itrev [] ys = ys" |
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    82
"itrev (x#xs) ys = itrev xs (x # ys)"
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    83
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    84
lemma itrev: "itrev xs ys = rev xs @ ys"
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    85
by(induction xs arbitrary: ys) auto
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    86
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    87
lemma itrev_Nil: "itrev xs [] = rev xs"
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    88
by(simp add: itrev)
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    89
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    90
time_fun itrev
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    91
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    92
lemma T_itrev: "T_itrev xs ys = length xs + 1"
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    93
by(induction xs arbitrary: ys) auto
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    94
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    95
time_fun tl
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    96
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    97
lemma T_tl: "T_tl xs = 0"
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    98
by (cases xs) simp_all
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
    99
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
   100
declare T_tl.simps[simp del]
8070e4578ece merged Reverse into Time_Funs
nipkow
parents: 81147
diff changeset
   101
73108
981a383610df HOL-Data_Structures: added Selection and time functions for list functions
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   102
end