doc-src/TutorialI/Recdef/examples.thy
changeset 11480 0fba0357c04c
parent 11458 09a6c44a48ea
child 11705 ac8ca15c556c
equal deleted inserted replaced
11479:697dcaaf478f 11480:0fba0357c04c
     7 *}
     7 *}
     8 
     8 
     9 consts fib :: "nat \<Rightarrow> nat";
     9 consts fib :: "nat \<Rightarrow> nat";
    10 recdef fib "measure(\<lambda>n. n)"
    10 recdef fib "measure(\<lambda>n. n)"
    11   "fib 0 = 0"
    11   "fib 0 = 0"
    12   "fib 1 = 1"
    12   "fib 1' = 1"
    13   "fib (Suc(Suc x)) = fib x + fib (Suc x)";
    13   "fib (Suc(Suc x)) = fib x + fib (Suc x)";
    14 
    14 
    15 text{*\noindent
    15 text{*\noindent
    16 \index{measure functions}%
    16 \index{measure functions}%
    17 The definition of @{term"fib"} is accompanied by a \textbf{measure function}
    17 The definition of @{term"fib"} is accompanied by a \textbf{measure function}