src/HOL/ex/LexOrds.thy
author bulwahn
Wed, 07 Feb 2007 13:05:28 +0100
changeset 22258 0967b03844b5
parent 21131 a447addc14af
child 22309 87ec1ca65312
permissions -rw-r--r--
changes in lexicographic_order termination tactic
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
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    12
use "~~/src/HOL/Tools/function_package/sum_tools.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    13
use "~~/src/HOL/Tools/function_package/fundef_common.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    14
use "~~/src/HOL/Tools/function_package/fundef_lib.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    15
use "~~/src/HOL/Tools/function_package/inductive_wrap.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    16
use "~~/src/HOL/Tools/function_package/context_tree.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    17
(* use "~~/src/HOL/Tools/function_package/fundef_tailrec.ML"*)
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    18
use "~~/src/HOL/Tools/function_package/fundef_prep.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    19
use "~~/src/HOL/Tools/function_package/fundef_proof.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    20
use "~~/src/HOL/Tools/function_package/termination.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    21
use "~~/src/HOL/Tools/function_package/mutual.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    22
use "~~/src/HOL/Tools/function_package/pattern_split.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    23
use "~~/src/HOL/Tools/function_package/auto_term.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    24
use "~~/src/HOL/Tools/function_package/fundef_package.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    25
use "~~/src/HOL/Tools/function_package/lexicographic_order.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    26
use "~~/src/HOL/Tools/function_package/fundef_datatype.ML"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    27
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    28
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    29
setup FundefPackage.setup
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    30
setup LexicographicOrder.setup
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    31
setup FundefDatatype.setup
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    32
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    33
lemmas [fundef_cong] = 
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    34
  let_cong if_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    35
  map_cong
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    36
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    37
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    38
subsection {* Trivial examples *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    39
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    40
fun f :: "nat \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    41
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    42
"f n = n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    43
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    44
fun g :: "nat \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    45
where 
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    46
  "g 0 = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    47
  "g (Suc n) = Suc (g n)"
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
subsection {* Examples on natural numbers *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    51
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    52
fun bin :: "(nat * nat) \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    53
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    54
  "bin (0, 0) = 1"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    55
  "bin (Suc n, 0) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    56
  "bin (0, Suc m) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    57
  "bin (Suc n, Suc m) = bin (n, m) + bin (Suc n, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    58
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    59
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    60
fun t :: "(nat * nat) \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    61
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    62
  "t (0,n) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    63
  "t (n,0) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    64
  "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
    65
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    66
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    67
function h :: "(nat * nat) * (nat * nat) \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    68
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    69
  "h ((0,0),(0,0)) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    70
  "h ((Suc z, y), (u,v)) = h((z, y), (u, v))" (* z is descending *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    71
  "h ((0, Suc y), (u,v)) = h((1, y), (u, v))" (* y is descending *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    72
  "h ((0,0), (Suc u, v)) = h((1, 1), (u, v))" (* u is descending *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    73
  "h ((0,0), (0, Suc v)) = h ((1,1), (1,v))" (*  v is descending *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    74
by (pat_completeness, auto)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    75
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    76
termination by lexicographic_order
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    77
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    78
lemma[simp]: "size x = x"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    79
apply (induct x) apply simp_all done
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    80
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    81
fun gcd2 :: "nat \<Rightarrow> nat \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    82
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    83
  "gcd2 x 0 = x"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    84
  "gcd2 0 y = y"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
    85
  "gcd2 (Suc x) (Suc y) = (if x < y then gcd2 (Suc x) (y - x)
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    86
                                    else gcd2 (x - y) (Suc y))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    87
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    88
termination by lexicographic_order
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    89
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    90
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    91
function ack :: "(nat * nat) \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    92
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    93
  "ack (0, m) = Suc m"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    94
  "ack (Suc n, 0) = ack(n, 1)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    95
  "ack (Suc n, Suc m) = ack (n, ack (Suc n, m))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    96
by pat_completeness auto
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    97
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    98
termination by lexicographic_order
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
    99
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   100
subsection {* Simple examples with other datatypes than nat, e.g. trees and lists *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   101
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   102
datatype tree = Node | Branch tree tree
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   103
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   104
fun g_tree :: "tree * tree \<Rightarrow> tree"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   105
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   106
  "g_tree (Node, Node) = Node"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   107
  "g_tree (Node, Branch a b) = Branch Node (g_tree (a,b))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   108
  "g_tree (Branch a b, Node) = Branch (g_tree (a,Node)) b"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   109
  "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
   110
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   111
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   112
termination by lexicographic_order
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   113
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   114
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   115
fun acklist :: "'a list * 'a list \<Rightarrow> 'a list"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   116
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   117
  "acklist ([], m) = ((hd m)#m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   118
|  "acklist (n#ns, []) = acklist (ns, [n])"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   119
|  "acklist ((n#ns), (m#ms)) = acklist (ns, acklist ((n#ns), ms))"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   120
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   121
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   122
subsection {* Examples with mutual recursion *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   123
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   124
fun evn od :: "nat \<Rightarrow> bool"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   125
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   126
  "evn 0 = True"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   127
| "od 0 = False"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   128
| "evn (Suc n) = od n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   129
| "od (Suc n) = evn n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   130
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   132
fun
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   133
 evn2 od2  :: "(nat * nat) \<Rightarrow> bool"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   134
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   135
  "evn2 (0, n) = True"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   136
  "evn2 (n, 0) = True"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   137
  "od2 (0, n) = False"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   138
  "od2 (n, 0) = False"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   139
  "evn2 (Suc n, Suc m) = od2 (Suc n, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   140
  "od2 (Suc n, Suc m) = evn2 (n, Suc m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   141
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   142
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   143
fun evn3 od3 :: "(nat * nat) \<Rightarrow> nat"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   144
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   145
  "evn3 (0,n) = n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   146
  "od3 (0,n) = n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   147
  "evn3 (n,0) = n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   148
  "od3 (n,0) = n"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   149
  "evn3 (Suc n, Suc m) = od3 (Suc m, n)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   150
  "od3 (Suc n, Suc m) = evn3 (Suc m, n) + od3(n, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   151
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   152
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   153
fun div3r0 div3r1 div3r2 :: "(nat * nat) \<Rightarrow> bool"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   154
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   155
  "div3r0 (0, 0) = True"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   156
  "div3r1 (0, 0) = False"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   157
  "div3r2 (0, 0) = False"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   158
  "div3r0 (0, Suc m) = div3r2 (0, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   159
  "div3r1 (0, Suc m) = div3r0 (0, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   160
  "div3r2 (0, Suc m) = div3r1 (0, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   161
  "div3r0 (Suc n, 0) = div3r2 (n, 0)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   162
  "div3r1 (Suc n, 0) = div3r0 (n, 0)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   163
  "div3r2 (Suc n, 0) = div3r1 (n, 0)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   164
  "div3r1 (Suc n, Suc m) = div3r2 (n, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   165
  "div3r2 (Suc n, Suc m) = div3r0 (n, m)"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   166
  "div3r0 (Suc n, Suc m) = div3r1 (n, m)"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   167
(*
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   168
lemma "i \<noteq> [] ==> sum_case (%x. size (fst x)) (%x. size (fst x)) (Inr (tl i, xa, i)) <  sum_case (%x. size (fst x)) (%x. size (fst x)) ( Inl (i, xa))"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   169
apply simp
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   170
done
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   171
*)
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   172
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   173
lemma "i \<noteq> [] \<Longrightarrow> size (tl i) < size i"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   174
apply simp
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   175
done
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   176
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   177
fun sizechange_f :: "'a list => 'a list => 'a list" and
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   178
sizechange_g :: "'a list => 'a list => 'a list => 'a list"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   179
where
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   180
"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
   181
"sizechange_g a b c = sizechange_f a (b @ c)"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   182
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   183
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   184
fun
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   185
  prod :: "nat => nat => nat => nat" and
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   186
eprod :: "nat => nat => nat => nat" and
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   187
oprod :: "nat => nat => nat => nat"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   188
where
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   189
"prod x y z = (if y mod 2 = 0 then eprod x y z else oprod x y z)"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   190
"oprod x y z = eprod x (y - 1) (z+x)"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   191
"eprod x y z = (if y=0 then z else prod (2*x) (y div 2) z)"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   192
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   193
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   194
function (sequential)
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   195
  pedal :: "nat => nat => nat => nat"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   196
and
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   197
  coast :: "nat => nat => nat => nat"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   198
where
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   199
  "pedal 0 m c = c"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   200
| "pedal n 0 c = c"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   201
| "pedal n m c =
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   202
     (if n < m then coast (n - 1) (m - 1) (c + m)
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   203
               else pedal (n - 1) m (c + m))"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   204
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   205
| "coast n m c =
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   206
     (if n < m then coast n (m - 1) (c + n)
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   207
               else pedal n m (c + n))"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   208
  by pat_completeness auto
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   209
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   210
lemma [simp]: "size (x::nat) = x"
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   211
by (induct x) auto 
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   212
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   213
termination apply lexicographic_order done
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   214
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   215
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   216
subsection {*Examples for an unprovable termination *}
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   217
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   218
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
   219
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   220
function noterm :: "(nat * nat) \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   221
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   222
  "noterm (a,b) = noterm(b,a)"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   223
by pat_completeness auto
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   224
(* termination by apply lexicographic_order*)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   225
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   226
function term_but_no_prove :: "nat * nat \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   227
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   228
  "term_but_no_prove (0,0) = 1"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   229
  "term_but_no_prove (0, Suc b) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   230
  "term_but_no_prove (Suc a, 0) = 0"
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   231
  "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
   232
by pat_completeness auto
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   233
(* termination by lexicographic_order *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   234
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   235
text{* The tactic distinguishes between N = not provable AND F = False *}
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   236
function no_proof :: "nat \<Rightarrow> nat"
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   237
where
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   238
  "no_proof m = no_proof (Suc m)"
22258
0967b03844b5 changes in lexicographic_order termination tactic
bulwahn
parents: 21131
diff changeset
   239
by pat_completeness auto
21131
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   240
(* termination by lexicographic_order *)
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   241
a447addc14af added lexicographic_order tactic
bulwahn
parents:
diff changeset
   242
end