src/HOL/Data_Structures/Base_FDS.thy
author wenzelm
Fri, 24 Nov 2023 15:58:24 +0100
changeset 79049 10b6add456d0
parent 70250 20d819b0a29d
permissions -rw-r--r--
support for Unix-domain sockets, using java.nio.channels.ServerSocketChannel;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70250
20d819b0a29d New version of tries
nipkow
parents: 67399
diff changeset
     1
section "Common Basis Theory"
20d819b0a29d New version of tries
nipkow
parents: 67399
diff changeset
     2
66491
nipkow
parents:
diff changeset
     3
theory Base_FDS
66672
75694b28ef08 updated imports;
wenzelm
parents: 66499
diff changeset
     4
imports "HOL-Library.Pattern_Aliases"
66491
nipkow
parents:
diff changeset
     5
begin
nipkow
parents:
diff changeset
     6
nipkow
parents:
diff changeset
     7
declare Let_def [simp]
nipkow
parents:
diff changeset
     8
nipkow
parents:
diff changeset
     9
text \<open>Lemma \<open>size_prod_measure\<close>, when declared with the \<open>measure_function\<close> attribute,
nipkow
parents:
diff changeset
    10
enables \<open>fun\<close> to prove termination of a larger class of functions automatically.
nipkow
parents:
diff changeset
    11
By default, \<open>fun\<close> only tries lexicographic combinations of the sizes of the parameters.
nipkow
parents:
diff changeset
    12
With \<open>size_prod_measure\<close> enabled it also tries measures based on the sum of the sizes
nipkow
parents:
diff changeset
    13
of different parameters.
nipkow
parents:
diff changeset
    14
nipkow
parents:
diff changeset
    15
To alert the reader whenever such a more subtle termination proof is taking place
66499
nipkow
parents: 66491
diff changeset
    16
the lemma is not enabled all the time but only when it is needed.
66491
nipkow
parents:
diff changeset
    17
\<close>
nipkow
parents:
diff changeset
    18
nipkow
parents:
diff changeset
    19
lemma size_prod_measure: 
nipkow
parents:
diff changeset
    20
  "is_measure f \<Longrightarrow> is_measure g \<Longrightarrow> is_measure (size_prod f g)"
nipkow
parents:
diff changeset
    21
by (rule is_measure_trivial)
nipkow
parents:
diff changeset
    22
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66672
diff changeset
    23
end