# HG changeset patch # User Manuel Eberl # Date 1565107446 -7200 # Node ID 8406a2c296e08c8f15d93226888c9047e2ae0dd1 # Parent b32d571f1190299ea6eb1f772f5ca8417744e999 Added Takeuchi function to HOL-ex diff -r b32d571f1190 -r 8406a2c296e0 src/HOL/ex/Functions.thy --- a/src/HOL/ex/Functions.thy Mon Aug 05 11:20:56 2019 +0200 +++ b/src/HOL/ex/Functions.thy Tue Aug 06 18:04:06 2019 +0200 @@ -96,6 +96,28 @@ by (induct n rule: f91.induct) auto +subsubsection \Here comes Takeuchi's function\ + +definition tak_m1 where "tak_m1 = (\(x,y,z). if x \ y then 0 else 1)" +definition tak_m2 where "tak_m2 = (\(x,y,z). nat (Max {x, y, z} - Min {x, y, z}))" +definition tak_m3 where "tak_m3 = (\(x,y,z). nat (x - Min {x, y, z}))" + +function tak :: "int \ int \ int \ int" where + "tak x y z = (if x \ y then y else tak (tak (x-1) y z) (tak (y-1) z x) (tak (z-1) x y))" + by auto + +lemma tak_pcorrect: + "tak_dom (x, y, z) \ tak x y z = (if x \ y then y else if y \ z then z else x)" + by (induction x y z rule: tak.pinduct) (auto simp: tak.psimps) + +termination + by (relation "tak_m1 <*mlex*> tak_m2 <*mlex*> tak_m3 <*mlex*> {}") + (auto simp: mlex_iff wf_mlex tak_pcorrect tak_m1_def tak_m2_def tak_m3_def min_def max_def) + +theorem tak_correct: "tak x y z = (if x \ y then y else if y \ z then z else x)" + by (induction x y z rule: tak.induct) auto + + subsection \More general patterns\ subsubsection \Overlapping patterns\ @@ -244,7 +266,6 @@ lemma "pred n = 0 \ n = 0 \ n = Suc 0" by (erule pred0E) metis+ - text \ Other expressions on the right-hand side also work, but whether the generated rule is useful depends on how well the simplifier can