| 70250 |      1 | section "Common Basis Theory"
 | 
|  |      2 | 
 | 
| 66491 |      3 | theory Base_FDS
 | 
| 66672 |      4 | imports "HOL-Library.Pattern_Aliases"
 | 
| 66491 |      5 | begin
 | 
|  |      6 | 
 | 
|  |      7 | declare Let_def [simp]
 | 
|  |      8 | 
 | 
|  |      9 | text \<open>Lemma \<open>size_prod_measure\<close>, when declared with the \<open>measure_function\<close> attribute,
 | 
|  |     10 | enables \<open>fun\<close> to prove termination of a larger class of functions automatically.
 | 
|  |     11 | By default, \<open>fun\<close> only tries lexicographic combinations of the sizes of the parameters.
 | 
|  |     12 | With \<open>size_prod_measure\<close> enabled it also tries measures based on the sum of the sizes
 | 
|  |     13 | of different parameters.
 | 
|  |     14 | 
 | 
|  |     15 | To alert the reader whenever such a more subtle termination proof is taking place
 | 
| 66499 |     16 | the lemma is not enabled all the time but only when it is needed.
 | 
| 66491 |     17 | \<close>
 | 
|  |     18 | 
 | 
|  |     19 | lemma size_prod_measure: 
 | 
|  |     20 |   "is_measure f \<Longrightarrow> is_measure g \<Longrightarrow> is_measure (size_prod f g)"
 | 
|  |     21 | by (rule is_measure_trivial)
 | 
|  |     22 | 
 | 
| 67399 |     23 | end
 |