src/HOL/Tools/Nitpick/nitpick_model.ML
changeset 38190 b02e204b613a
parent 38189 a493dc2179a3
child 38207 792b78e355e7
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML	Tue Aug 03 21:37:12 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML	Wed Aug 04 10:39:35 2010 +0200
@@ -170,7 +170,7 @@
                 Type (@{type_name fun}, [T1, Type (@{type_name fun}, [T2, _])]))
          $ t1 $ t2) =
     let val Ts = map uniterize_unarize_unbox_etc_type [T1, T2] in
-      Const (@{const_name Pair}, Ts ---> Type (@{type_name Product_Type.prod}, Ts))
+      Const (@{const_name Pair}, Ts ---> Type (@{type_name prod}, Ts))
       $ unarize_unbox_etc_term t1 $ unarize_unbox_etc_term t2
     end
   | unarize_unbox_etc_term (Const (s, T)) =
@@ -185,27 +185,27 @@
   | unarize_unbox_etc_term (Abs (s, T, t')) =
     Abs (s, uniterize_unarize_unbox_etc_type T, unarize_unbox_etc_term t')
 
-fun factor_out_types (T1 as Type (@{type_name Product_Type.prod}, [T11, T12]))
-                     (T2 as Type (@{type_name Product_Type.prod}, [T21, T22])) =
+fun factor_out_types (T1 as Type (@{type_name prod}, [T11, T12]))
+                     (T2 as Type (@{type_name prod}, [T21, T22])) =
     let val (n1, n2) = pairself num_factors_in_type (T11, T21) in
       if n1 = n2 then
         let
           val ((T11', opt_T12'), (T21', opt_T22')) = factor_out_types T12 T22
         in
-          ((Type (@{type_name Product_Type.prod}, [T11, T11']), opt_T12'),
-           (Type (@{type_name Product_Type.prod}, [T21, T21']), opt_T22'))
+          ((Type (@{type_name prod}, [T11, T11']), opt_T12'),
+           (Type (@{type_name prod}, [T21, T21']), opt_T22'))
         end
       else if n1 < n2 then
         case factor_out_types T1 T21 of
           (p1, (T21', NONE)) => (p1, (T21', SOME T22))
         | (p1, (T21', SOME T22')) =>
-          (p1, (T21', SOME (Type (@{type_name Product_Type.prod}, [T22', T22]))))
+          (p1, (T21', SOME (Type (@{type_name prod}, [T22', T22]))))
       else
         swap (factor_out_types T2 T1)
     end
-  | factor_out_types (Type (@{type_name Product_Type.prod}, [T11, T12])) T2 =
+  | factor_out_types (Type (@{type_name prod}, [T11, T12])) T2 =
     ((T11, SOME T12), (T2, NONE))
-  | factor_out_types T1 (Type (@{type_name Product_Type.prod}, [T21, T22])) =
+  | factor_out_types T1 (Type (@{type_name prod}, [T21, T22])) =
     ((T1, NONE), (T21, SOME T22))
   | factor_out_types T1 T2 = ((T1, NONE), (T2, NONE))
 
@@ -239,7 +239,7 @@
     val (ps1, ps2) = pairself HOLogic.flat_tupleT_paths (T1, T2)
     val (ts1, ts2) = t |> HOLogic.strip_ptuple ps |> chop cut
   in (HOLogic.mk_ptuple ps1 T1 ts1, HOLogic.mk_ptuple ps2 T2 ts2) end
-fun pair_up (Type (@{type_name Product_Type.prod}, [T1', T2']))
+fun pair_up (Type (@{type_name prod}, [T1', T2']))
             (t1 as Const (@{const_name Pair},
                           Type (@{type_name fun},
                                 [_, Type (@{type_name fun}, [_, T1])]))
@@ -287,8 +287,8 @@
       and do_term (Type (@{type_name fun}, [T1', T2']))
                   (Type (@{type_name fun}, [T1, T2])) t =
           do_fun T1' T2' T1 T2 t
-        | do_term (T' as Type (@{type_name Product_Type.prod}, Ts' as [T1', T2']))
-                  (Type (@{type_name Product_Type.prod}, [T1, T2]))
+        | do_term (T' as Type (@{type_name prod}, Ts' as [T1', T2']))
+                  (Type (@{type_name prod}, [T1, T2]))
                   (Const (@{const_name Pair}, _) $ t1 $ t2) =
           Const (@{const_name Pair}, Ts' ---> T')
           $ do_term T1' T1 t1 $ do_term T2' T2 t2
@@ -303,7 +303,7 @@
   | truth_const_sort_key @{const False} = "2"
   | truth_const_sort_key _ = "1"
 
-fun mk_tuple (Type (@{type_name Product_Type.prod}, [T1, T2])) ts =
+fun mk_tuple (Type (@{type_name prod}, [T1, T2])) ts =
     HOLogic.mk_prod (mk_tuple T1 ts,
         mk_tuple T2 (List.drop (ts, length (HOLogic.flatten_tupleT T1))))
   | mk_tuple _ (t :: _) = t
@@ -463,7 +463,7 @@
                             signed_string_of_int j ^ " for " ^
                             string_for_rep (Vect (k1, Atom (k2, 0))))
         end
-      | term_for_atom seen (Type (@{type_name Product_Type.prod}, [T1, T2])) _ j k =
+      | term_for_atom seen (Type (@{type_name prod}, [T1, T2])) _ j k =
         let
           val k1 = card_of_type card_assigns T1
           val k2 = k div k1
@@ -476,7 +476,6 @@
         HOLogic.mk_Trueprop (term_for_atom seen bool_T bool_T j k)
       | term_for_atom _ @{typ bool} _ j _ =
         if j = 0 then @{const False} else @{const True}
-      | term_for_atom _ @{typ unit} _ _ _ = @{const Unity}
       | term_for_atom seen T _ j k =
         if T = nat_T andalso is_standard_datatype thy stds nat_T then
           HOLogic.mk_number nat_T j
@@ -588,11 +587,10 @@
                (map (fn j => term_for_atom seen T1 T1 j k) (index_seq 0 k))
                (map (term_for_rep true seen T2 T2 R o single)
                     (batch_list (arity_of_rep R) js))
-    and term_for_rep _ seen T T' Unit [[]] = term_for_atom seen T T' 0 1
-      | term_for_rep _ seen T T' (R as Atom (k, j0)) [[j]] =
+    and term_for_rep _ seen T T' (R as Atom (k, j0)) [[j]] =
         if j >= j0 andalso j < j0 + k then term_for_atom seen T T' (j - j0) k
         else raise REP ("Nitpick_Model.reconstruct_term.term_for_rep", [R])
-      | term_for_rep _ seen (Type (@{type_name Product_Type.prod}, [T1, T2])) _
+      | term_for_rep _ seen (Type (@{type_name prod}, [T1, T2])) _
                      (Struct [R1, R2]) [js] =
         let
           val arity1 = arity_of_rep R1