coercion-invariant arguments at work
authortraytel
Fri, 01 Mar 2013 22:15:31 +0100
changeset 51320 c1cb872ccb2b
parent 51319 4a92178011e7
child 51321 75682dfff630
coercion-invariant arguments at work
src/HOL/ex/Coercion_Examples.thy
--- a/src/HOL/ex/Coercion_Examples.thy	Fri Mar 01 22:15:31 2013 +0100
+++ b/src/HOL/ex/Coercion_Examples.thy	Fri Mar 01 22:15:31 2013 +0100
@@ -178,4 +178,22 @@
 
 term "ys=[[[[[1::nat]]]]]"
 
+consts
+  case_nil :: "'a \<Rightarrow> 'b"
+  case_cons :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+  case_abs :: "('c \<Rightarrow> 'b) \<Rightarrow> 'b"
+  case_elem :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"
+
+declare [[coercion_args case_cons - -]]
+declare [[coercion_args case_abs -]]
+declare [[coercion_args case_elem - +]]
+
+term "case_cons (case_abs (\<lambda>n. case_abs (\<lambda>is. case_elem (((n::nat),(is::int list))) (n#is)))) case_nil"
+
+consts n :: nat m :: nat
+term "- (n + m)"
+declare [[coercion_args uminus -]]
+declare [[coercion_args plus + +]]
+term "- (n + m)"
+ 
 end