src/HOL/Data_Structures/Cmp.thy
changeset 61581 00d9682e8dd7
child 61640 44c9198f210c
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Data_Structures/Cmp.thy	Thu Nov 05 08:27:14 2015 +0100
@@ -0,0 +1,21 @@
+(* Author: Tobias Nipkow *)
+
+section {* Three-Way Comparison *}
+
+theory Cmp
+imports Main
+begin
+
+datatype cmp = LT | EQ | GT
+
+class cmp = linorder +
+fixes cmp :: "'a \<Rightarrow> 'a \<Rightarrow> cmp"
+assumes LT[simp]: "cmp x y = LT \<longleftrightarrow> x < y"
+assumes EQ[simp]: "cmp x y = EQ \<longleftrightarrow> x = y"
+assumes GT[simp]: "cmp x y = GT \<longleftrightarrow> x > y"
+
+lemma case_cmp_if[simp]: "(case c of EQ \<Rightarrow> e | LT \<Rightarrow> l | GT \<Rightarrow> g) =
+  (if c = LT then l else if c = GT then g else e)"
+by(simp split: cmp.split)
+
+end