shortened definitions by adding some termify constants
authorbulwahn
Fri, 20 Jan 2012 09:28:52 +0100
changeset 46307 ec8f975c059b
parent 46306 940ddb42c998
child 46308 e5abbec2697a
shortened definitions by adding some termify constants
src/HOL/Quickcheck_Exhaustive.thy
--- a/src/HOL/Quickcheck_Exhaustive.thy	Fri Jan 20 09:28:51 2012 +0100
+++ b/src/HOL/Quickcheck_Exhaustive.thy	Fri Jan 20 09:28:52 2012 +0100
@@ -119,17 +119,13 @@
 
 end
 
+definition (in term_syntax) [code_unfold]: "valtermify_pair x y = Code_Evaluation.valtermify (Pair :: 'a :: typerep => 'b :: typerep => 'a * 'b) {\<cdot>} x {\<cdot>} y"
+
 instantiation prod :: (full_exhaustive, full_exhaustive) full_exhaustive
 begin
 
 definition
-  "full_exhaustive f d = full_exhaustive (%(x, t1). full_exhaustive (%(y, t2). f ((x, y),
-    %u. let T1 = (Typerep.typerep (TYPE('a)));
-            T2 = (Typerep.typerep (TYPE('b)))
-    in Code_Evaluation.App (Code_Evaluation.App (
-      Code_Evaluation.Const (STR ''Product_Type.Pair'') 
-      (Typerep.Typerep (STR ''fun'') [T1, Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''Product_Type.prod'') [T1, T2]]]))
-      (t1 ())) (t2 ()))) d) d"
+  "full_exhaustive f d = full_exhaustive (%x. full_exhaustive (%y. f (valtermify_pair x y)) d) d"
 
 instance ..
 
@@ -140,7 +136,7 @@
 
 fun exhaustive_set
 where
-  "exhaustive_set f i = (if i = 0 then None else (f {} orelse exhaustive_set (%A. f A orelse Quickcheck_Exhaustive.exhaustive (%x. if x \<in> A then None else f (insert x A)) (i - 1)) (i - 1)))"
+  "exhaustive_set f i = (if i = 0 then None else (f {} orelse exhaustive_set (%A. f A orelse exhaustive (%x. if x \<in> A then None else f (insert x A)) (i - 1)) (i - 1)))"
 
 instance ..
 
@@ -178,22 +174,19 @@
 
 end
 
+definition [code_unfold]: "valtermify_absdummy = (%(v, t). (%_::'a. v, %u::unit. Code_Evaluation.Abs (STR ''x'') (Typerep.typerep TYPE('a::typerep)) (t ())))"
+
+definition (in term_syntax) [code_unfold]: "valtermify_fun_upd g a b = Code_Evaluation.valtermify (fun_upd :: ('a :: typerep => 'b :: typerep) => 'a => 'b => 'a => 'b) {\<cdot>} g {\<cdot>} a {\<cdot>} b"
+
 instantiation "fun" :: ("{equal, full_exhaustive}", full_exhaustive) full_exhaustive
 begin
 
 fun full_exhaustive_fun' :: "(('a => 'b) * (unit => term) => (bool * term list) option) => code_numeral => code_numeral => (bool * term list) option"
 where
-  "full_exhaustive_fun' f i d = (full_exhaustive (%(b, t). f (%_. b, %_. Code_Evaluation.Abs (STR ''x'') (Typerep.typerep TYPE('a)) (t ()))) d)
+  "full_exhaustive_fun' f i d = (full_exhaustive (%v. f (valtermify_absdummy v)) d)
    orelse (if i > 1 then
-     full_exhaustive_fun' (%(g, gt). full_exhaustive (%(a, at). full_exhaustive (%(b, bt).
-       f (g(a := b),
-         (%_. let A = (Typerep.typerep (TYPE('a)));
-                  B = (Typerep.typerep (TYPE('b)));
-                  fun = (%T U. Typerep.Typerep (STR ''fun'') [T, U])
-              in
-                Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.App
-                  (Code_Evaluation.Const (STR ''Fun.fun_upd'') (fun (fun A B) (fun A (fun B (fun A B)))))
-                (gt ())) (at ())) (bt ())))) d) d) (i - 1) d else None)"
+     full_exhaustive_fun' (%g. full_exhaustive (%a. full_exhaustive (%b.
+       f (valtermify_fun_upd g a b)) d) d) (i - 1) d else None)"
 
 definition full_exhaustive_fun :: "(('a => 'b) * (unit => term) => (bool * term list) option) => code_numeral => (bool * term list) option"
 where
@@ -214,6 +207,8 @@
   "check_all_n_lists f n =
      (if n = 0 then f ([], (%_. [])) else check_all (%(x, xt). check_all_n_lists (%(xs, xst). f ((x # xs), (%_. (xt () # xst ())))) (n - 1)))"
 
+definition (in term_syntax) [code_unfold]: "termify_fun_upd g a b = (Code_Evaluation.termify (fun_upd :: ('a :: typerep => 'b :: typerep) => 'a => 'b => 'a => 'b) <\<cdot>> g <\<cdot>> a <\<cdot>> b)"
+
 definition mk_map_term :: " (unit \<Rightarrow> typerep) \<Rightarrow> (unit \<Rightarrow> typerep) \<Rightarrow> (unit \<Rightarrow> term list) \<Rightarrow> (unit \<Rightarrow> term list) \<Rightarrow> unit \<Rightarrow> term"
 where
   "mk_map_term T1 T2 domm rng =
@@ -307,48 +302,31 @@
 
 end
 
+definition (in term_syntax) [code_unfold]: "termify_pair x y = Code_Evaluation.termify (Pair :: 'a :: typerep => 'b :: typerep => 'a * 'b) <\<cdot>> x <\<cdot>> y"
 
 instantiation prod :: (check_all, check_all) check_all
 begin
 
 definition
-  "check_all f = check_all (%(x, t1). check_all (%(y, t2). f ((x, y),
-    %u. let T1 = (Typerep.typerep (TYPE('a)));
-            T2 = (Typerep.typerep (TYPE('b)))
-    in Code_Evaluation.App (Code_Evaluation.App (
-      Code_Evaluation.Const (STR ''Product_Type.Pair'') 
-      (Typerep.Typerep (STR ''fun'') [T1, Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''Product_Type.prod'') [T1, T2]]]))
-      (t1 ())) (t2 ()))))"
+  "check_all f = check_all (%x. check_all (%y. f (valtermify_pair x y)))"
 
 definition enum_term_of_prod :: "('a * 'b) itself => unit => term list"
 where
-  "enum_term_of_prod = (%_ _. map (%(x, y).
-       let T1 = (Typerep.typerep (TYPE('a)));
-           T2 = (Typerep.typerep (TYPE('b)))
-       in Code_Evaluation.App (Code_Evaluation.App (
-         Code_Evaluation.Const (STR ''Product_Type.Pair'') 
-           (Typerep.Typerep (STR ''fun'') [T1, Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''Product_Type.prod'') [T1, T2]]])) x) y)
-     (Enum.product (enum_term_of (TYPE('a)) ()) (enum_term_of (TYPE('b)) ())))  "
+  "enum_term_of_prod = (%_ _. map (%(x, y). termify_pair TYPE('a) TYPE('b) x y)
+     (Enum.product (enum_term_of (TYPE('a)) ()) (enum_term_of (TYPE('b)) ())))"
 
 instance ..
 
 end
 
+definition (in term_syntax) [code_unfold]: "valtermify_Inl x = Code_Evaluation.valtermify (Inl :: 'a :: typerep => 'a + 'b :: typerep) {\<cdot>} x"
+definition (in term_syntax) [code_unfold]: "valtermify_Inr x = Code_Evaluation.valtermify (Inr :: 'b :: typerep => 'a ::typerep + 'b) {\<cdot>} x"
 
 instantiation sum :: (check_all, check_all) check_all
 begin
 
 definition
-  "check_all f = (case check_all (%(a, t). f (Inl a, %_. 
-     let T1 = (Typerep.typerep (TYPE('a)));
-         T2 = (Typerep.typerep (TYPE('b)))
-       in Code_Evaluation.App (Code_Evaluation.Const (STR ''Sum_Type.Inl'') 
-           (Typerep.Typerep (STR ''fun'') [T1, Typerep.Typerep (STR ''Sum_Type.sum'') [T1, T2]])) (t ()))) of Some x' => Some x'
-             | None => check_all (%(b, t). f (Inr b, %_. let
-                 T1 = (Typerep.typerep (TYPE('a)));
-                 T2 = (Typerep.typerep (TYPE('b)))
-               in Code_Evaluation.App (Code_Evaluation.Const (STR ''Sum_Type.Inr'') 
-                 (Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''Sum_Type.sum'') [T1, T2]])) (t ()))))"
+  "check_all f = check_all (%a. f (valtermify_Inl a)) orelse check_all (%b. f (valtermify_Inr b))"
 
 definition enum_term_of_sum :: "('a + 'b) itself => unit => term list"
 where
@@ -598,7 +576,9 @@
   exhaustive'_def
   exhaustive_code_numeral'_def
 
-hide_const valterm_emptyset valtermify_insert term_emptyset termify_insert setify
+hide_const valtermify_absdummy valtermify_fun_upd valterm_emptyset valtermify_insert valtermify_pair
+  valtermify_Inl valtermify_Inr
+  termify_fun_upd term_emptyset termify_insert termify_pair setify
 
 hide_const (open)
   exhaustive full_exhaustive exhaustive' exhaustive_code_numeral' full_exhaustive_code_numeral'