src/HOL/ex/LexOrds.thy
author aspinall
Sat, 17 Feb 2007 17:18:47 +0100
changeset 22334 4c96d3370186
parent 22309 87ec1ca65312
child 27742 df552e6027cf
permissions -rw-r--r--
Clarify comment
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     1
(* Title:       HOL/ex/LexOrds.thy
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     2
   ID:
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     3
   Author:      Lukas Bulwahn, TU Muenchen
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     4
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     5
Examples for functions whose termination is proven by lexicographic order.
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     6
*)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     7
 
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     8
theory LexOrds
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
     9
imports Main
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    10
begin
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    11
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    12
subsection {* Trivial examples *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    13
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    14
fun identity :: "nat \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    15
where
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    16
"identity n = n"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    17
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    18
fun yaSuc :: "nat \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    19
where 
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    20
  "yaSuc 0 = 0"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    21
  "yaSuc (Suc n) = Suc (yaSuc n)"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    22
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    23
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    24
subsection {* Examples on natural numbers *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    25
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    26
fun bin :: "(nat * nat) \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    27
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    28
  "bin (0, 0) = 1"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    29
  "bin (Suc n, 0) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    30
  "bin (0, Suc m) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    31
  "bin (Suc n, Suc m) = bin (n, m) + bin (Suc n, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    32
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    33
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    34
fun t :: "(nat * nat) \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    35
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    36
  "t (0,n) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    37
  "t (n,0) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    38
  "t (Suc n, Suc m) = (if (n mod 2 = 0) then (t (Suc n, m)) else (t (n, Suc m)))" 
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    39
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    40
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    41
fun k :: "(nat * nat) * (nat * nat) \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    42
where
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    43
  "k ((0,0),(0,0)) = 0"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    44
  "k ((Suc z, y), (u,v)) = k((z, y), (u, v))" (* z is descending *)
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    45
  "k ((0, Suc y), (u,v)) = k((1, y), (u, v))" (* y is descending *)
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    46
  "k ((0,0), (Suc u, v)) = k((1, 1), (u, v))" (* u is descending *)
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    47
  "k ((0,0), (0, Suc v)) = k((1,1), (1,v))"   (* v is descending *)
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    48
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    49
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    50
fun gcd2 :: "nat \<Rightarrow> nat \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    51
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    52
  "gcd2 x 0 = x"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    53
  "gcd2 0 y = y"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    54
  "gcd2 (Suc x) (Suc y) = (if x < y then gcd2 (Suc x) (y - x)
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    55
                                    else gcd2 (x - y) (Suc y))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    56
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    57
fun ack :: "(nat * nat) \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    58
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    59
  "ack (0, m) = Suc m"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    60
  "ack (Suc n, 0) = ack(n, 1)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    61
  "ack (Suc n, Suc m) = ack (n, ack (Suc n, m))"
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    62
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    63
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    64
fun greedy :: "nat * nat * nat * nat * nat => nat"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    65
where
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    66
  "greedy (Suc a, Suc b, Suc c, Suc d, Suc e) =
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    67
  (if (a < 10) then greedy (Suc a, Suc b, c, d + 2, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    68
  (if (a < 20) then greedy (Suc a, b, Suc c, d, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    69
  (if (a < 30) then greedy (Suc a, b, Suc c, d, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    70
  (if (a < 40) then greedy (Suc a, b, Suc c, d, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    71
  (if (a < 50) then greedy (Suc a, b, Suc c, d, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    72
  (if (a < 60) then greedy (a, Suc b, Suc c, d, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    73
  (if (a < 70) then greedy (a, Suc b, Suc c, d, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    74
  (if (a < 80) then greedy (a, Suc b, Suc c, d, Suc e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    75
  (if (a < 90) then greedy (Suc a, Suc b, Suc c, d, e) else
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    76
  greedy (Suc a, Suc b, Suc c, d, e))))))))))"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    77
  "greedy (a, b, c, d, e) = 0"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    78
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    79
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    80
fun blowup :: "nat => nat => nat => nat => nat => nat => nat => nat => nat => nat"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    81
where
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    82
  "blowup 0 0 0 0 0 0 0 0 0 = 0"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    83
  "blowup 0 0 0 0 0 0 0 0 (Suc i) = Suc (blowup i i i i i i i i i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    84
  "blowup 0 0 0 0 0 0 0 (Suc h) i = Suc (blowup h h h h h h h h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    85
  "blowup 0 0 0 0 0 0 (Suc g) h i = Suc (blowup g g g g g g g h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    86
  "blowup 0 0 0 0 0 (Suc f) g h i = Suc (blowup f f f f f f g h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    87
  "blowup 0 0 0 0 (Suc e) f g h i = Suc (blowup e e e e e f g h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    88
  "blowup 0 0 0 (Suc d) e f g h i = Suc (blowup d d d d e f g h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    89
  "blowup 0 0 (Suc c) d e f g h i = Suc (blowup c c c d e f g h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    90
  "blowup 0 (Suc b) c d e f g h i = Suc (blowup b b c d e f g h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    91
  "blowup (Suc a) b c d e f g h i = Suc (blowup a b c d e f g h i)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    92
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
    93
  
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    94
subsection {* Simple examples with other datatypes than nat, e.g. trees and lists *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    95
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    96
datatype tree = Node | Branch tree tree
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    97
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    98
fun g_tree :: "tree * tree \<Rightarrow> tree"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    99
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   100
  "g_tree (Node, Node) = Node"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   101
  "g_tree (Node, Branch a b) = Branch Node (g_tree (a,b))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   102
  "g_tree (Branch a b, Node) = Branch (g_tree (a,Node)) b"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   103
  "g_tree (Branch a b, Branch c d) = Branch (g_tree (a,c)) (g_tree (b,d))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   104
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   105
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   106
fun acklist :: "'a list * 'a list \<Rightarrow> 'a list"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   107
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   108
  "acklist ([], m) = ((hd m)#m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   109
|  "acklist (n#ns, []) = acklist (ns, [n])"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   110
|  "acklist ((n#ns), (m#ms)) = acklist (ns, acklist ((n#ns), ms))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   111
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   112
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   113
subsection {* Examples with mutual recursion *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   114
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   115
fun evn od :: "nat \<Rightarrow> bool"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   116
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   117
  "evn 0 = True"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   118
| "od 0 = False"
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   119
| "evn (Suc n) = od (Suc n)"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   120
| "od (Suc n) = evn n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   121
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   122
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   123
fun sizechange_f :: "'a list => 'a list => 'a list" and
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   124
sizechange_g :: "'a list => 'a list => 'a list => 'a list"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   125
where
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   126
"sizechange_f i x = (if i=[] then x else sizechange_g (tl i) x i)"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   127
"sizechange_g a b c = sizechange_f a (b @ c)"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   128
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   129
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   130
fun
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   131
  prod :: "nat => nat => nat => nat" and
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   132
  eprod :: "nat => nat => nat => nat" and
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   133
  oprod :: "nat => nat => nat => nat"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   134
where
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   135
  "prod x y z = (if y mod 2 = 0 then eprod x y z else oprod x y z)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   136
  "oprod x y z = eprod x (y - 1) (z+x)"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   137
  "eprod x y z = (if y=0 then z else prod (2*x) (y div 2) z)"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   138
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   139
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   140
fun
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   141
  pedal :: "nat => nat => nat => nat"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   142
and
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   143
  coast :: "nat => nat => nat => nat"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   144
where
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   145
  "pedal 0 m c = c"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   146
| "pedal n 0 c = c"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   147
| "pedal n m c =
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   148
     (if n < m then coast (n - 1) (m - 1) (c + m)
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   149
               else pedal (n - 1) m (c + m))"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   150
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   151
| "coast n m c =
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   152
     (if n < m then coast n (m - 1) (c + n)
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   153
               else pedal n m (c + n))"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   154
22309
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   155
fun ack1 :: "nat => nat => nat"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   156
  and ack2 :: "nat => nat => nat"
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   157
  where
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   158
  "ack1 0 m = m+1" |
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   159
  "ack1 (Suc n) m = ack2 n m" |
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   160
  "ack2 n 0 = ack1 n 1" |
87ec1ca65312 improved lexicographic order termination tactic
bulwahn
parents: 22258
diff changeset
   161
  "ack2 n (Suc m) = ack1 n (ack2 n (Suc m))"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   162
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   163
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   164
subsection {*Examples for an unprovable termination *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   165
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   166
text {* If termination cannot be proven, the tactic gives further information about unprovable subgoals on the arguments *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   167
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   168
function noterm :: "(nat * nat) \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   169
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   170
  "noterm (a,b) = noterm(b,a)"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   171
by pat_completeness auto
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   172
(* termination by apply lexicographic_order*)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   173
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   174
function term_but_no_prove :: "nat * nat \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   175
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   176
  "term_but_no_prove (0,0) = 1"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   177
  "term_but_no_prove (0, Suc b) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   178
  "term_but_no_prove (Suc a, 0) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   179
  "term_but_no_prove (Suc a, Suc b) = term_but_no_prove (b, a)"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   180
by pat_completeness auto
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   181
(* termination by lexicographic_order *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   182
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   183
text{* The tactic distinguishes between N = not provable AND F = False *}
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   184
function no_proof :: "nat \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   185
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   186
  "no_proof m = no_proof (Suc m)"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   187
by pat_completeness auto
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   188
(* termination by lexicographic_order *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   189
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   190
end