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