src/HOL/ex/Recdefs.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 14953 27decf8d40ff
child 37456 0a1cc2675958
permissions -rw-r--r--
migrated theory headers to new format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6481
dbf2d9b3d6c8 recdef requires theory Recdef;
wenzelm
parents: 6455
diff changeset
     1
(*  Title:      HOL/ex/Recdefs.thy
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
     3
    Author:     Konrad Slind and Lawrence C Paulson
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
     5
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
     6
Examples of recdef definitions.  Most, but not all, are handled automatically.
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
     7
*)
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
     8
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
     9
header {* Examples of recdef definitions *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    10
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14953
diff changeset
    11
theory Recdefs imports Main begin
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    12
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    13
consts fact :: "nat => nat"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    14
recdef fact  less_than
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    15
  "fact x = (if x = 0 then 1 else x * fact (x - 1))"
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    16
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    17
consts Fact :: "nat => nat"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    18
recdef Fact  less_than
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    19
  "Fact 0 = 1"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    20
  "Fact (Suc x) = Fact x * Suc x"
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    21
14953
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    22
consts fib :: "int => int"
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    23
recdef fib  "measure nat"
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    24
  eqn:  "fib n = (if n < 1 then 0
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    25
                  else if n=1 then 1
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    26
                  else fib(n - 2) + fib(n - 1))";
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    27
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    28
lemma "fib 7 = 13"
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    29
by simp
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    30
27decf8d40ff new fib example
paulson
parents: 14244
diff changeset
    31
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    32
consts map2 :: "('a => 'b => 'c) * 'a list * 'b list => 'c list"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    33
recdef map2  "measure(\<lambda>(f, l1, l2). size l1)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    34
  "map2 (f, [], [])  = []"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    35
  "map2 (f, h # t, []) = []"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    36
  "map2 (f, h1 # t1, h2 # t2) = f h1 h2 # map2 (f, t1, t2)"
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    37
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    38
consts finiteRchain :: "('a => 'a => bool) * 'a list => bool"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    39
recdef finiteRchain  "measure (\<lambda>(R, l). size l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    40
  "finiteRchain(R,  []) = True"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    41
  "finiteRchain(R, [x]) = True"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    42
  "finiteRchain(R, x # y # rst) = (R x y \<and> finiteRchain (R, y # rst))"
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    43
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    44
text {* Not handled automatically: too complicated. *}
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    45
consts variant :: "nat * nat list => nat"
14244
f58598341d30 InductiveInvariant_examples illustrates advanced recursive function definitions
paulson
parents: 12023
diff changeset
    46
recdef (permissive) variant "measure (\<lambda>(n,ns). size (filter (\<lambda>y. n \<le> y) ns))"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    47
  "variant (x, L) = (if x mem L then variant (Suc x, L) else x)"
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    48
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    49
consts gcd :: "nat * nat => nat"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    50
recdef gcd  "measure (\<lambda>(x, y). x + y)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    51
  "gcd (0, y) = y"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    52
  "gcd (Suc x, 0) = Suc x"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    53
  "gcd (Suc x, Suc y) =
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    54
    (if y \<le> x then gcd (x - y, Suc y) else gcd (Suc x, y - x))"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    55
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    56
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    57
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    58
  \medskip The silly @{term g} function: example of nested recursion.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    59
  Not handled automatically.  In fact, @{term g} is the zero constant
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    60
  function.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    61
 *}
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    62
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    63
consts g :: "nat => nat"
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11024
diff changeset
    64
recdef (permissive) g  less_than
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    65
  "g 0 = 0"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    66
  "g (Suc x) = g (g x)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    67
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    68
lemma g_terminates: "g x < Suc x"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    69
  apply (induct x rule: g.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    70
   apply (auto simp add: g.simps)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    71
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    72
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    73
lemma g_zero: "g x = 0"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    74
  apply (induct x rule: g.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    75
   apply (simp_all add: g.simps g_terminates)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    76
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    77
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    78
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    79
consts Div :: "nat * nat => nat * nat"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    80
recdef Div  "measure fst"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    81
  "Div (0, x) = (0, 0)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    82
  "Div (Suc x, y) =
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    83
    (let (q, r) = Div (x, y)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    84
    in if y \<le> Suc r then (Suc q, 0) else (q, Suc r))"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    85
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    86
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    87
  \medskip Not handled automatically.  Should be the predecessor
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    88
  function, but there is an unnecessary "looping" recursive call in
12023
wenzelm
parents: 11701
diff changeset
    89
  @{text "k 1"}.
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    90
*}
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
    91
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    92
consts k :: "nat => nat"
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11626
diff changeset
    93
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11024
diff changeset
    94
recdef (permissive) k  less_than
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    95
  "k 0 = 0"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    96
  "k (Suc n) =
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    97
   (let x = k 1
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11626
diff changeset
    98
    in if False then k (Suc 1) else n)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
    99
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   100
consts part :: "('a => bool) * 'a list * 'a list * 'a list => 'a list * 'a list"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   101
recdef part  "measure (\<lambda>(P, l, l1, l2). size l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   102
  "part (P, [], l1, l2) = (l1, l2)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   103
  "part (P, h # rst, l1, l2) =
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   104
    (if P h then part (P, rst, h # l1, l2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   105
    else part (P, rst, l1, h # l2))"
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
   106
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   107
consts fqsort :: "('a => 'a => bool) * 'a list => 'a list"
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11024
diff changeset
   108
recdef (permissive) fqsort  "measure (size o snd)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   109
  "fqsort (ord, []) = []"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   110
  "fqsort (ord, x # rst) =
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   111
  (let (less, more) = part ((\<lambda>y. ord y x), rst, ([], []))
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   112
  in fqsort (ord, less) @ [x] @ fqsort (ord, more))"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   113
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   114
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   115
  \medskip Silly example which demonstrates the occasional need for
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   116
  additional congruence rules (here: @{thm [source] map_cong}).  If
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   117
  the congruence rule is removed, an unprovable termination condition
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   118
  is generated!  Termination not proved automatically.  TFL requires
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   119
  @{term [source] "\<lambda>x. mapf x"} instead of @{term [source] mapf}.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   120
*}
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
   121
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   122
consts mapf :: "nat => nat list"
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11024
diff changeset
   123
recdef (permissive) mapf  "measure (\<lambda>m. m)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   124
  "mapf 0 = []"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   125
  "mapf (Suc n) = concat (map (\<lambda>x. mapf x) (replicate n n))"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   126
  (hints cong: map_cong)
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
   127
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   128
recdef_tc mapf_tc: mapf
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   129
  apply (rule allI)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   130
  apply (case_tac "n = 0")
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   131
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   132
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   133
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   134
text {* Removing the termination condition from the generated thms: *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   135
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   136
lemma "mapf (Suc n) = concat (map mapf (replicate n n))"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   137
  apply (simp add: mapf.simps mapf_tc)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   138
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   139
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8522
diff changeset
   140
lemmas mapf_induct = mapf.induct [OF mapf_tc]
5124
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
   141
1ce3cccfacdb stepping stones: Recdef, Main;
wenzelm
parents:
diff changeset
   142
end