src/HOL/Tools/ATP/atp_util.ML
changeset 44284 717880e98e6b
parent 44012 37e1431cc213
child 44434 ae612a423dad
     1.1 --- a/src/HOL/Tools/ATP/atp_util.ML	Thu Jun 16 13:50:35 2011 +0200
     1.2 +++ b/src/HOL/Tools/ATP/atp_util.ML	Thu Jun 16 13:50:35 2011 +0200
     1.3 @@ -22,7 +22,7 @@
     1.4      Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
     1.5      -> typ
     1.6    val is_type_surely_finite : Proof.context -> typ -> bool
     1.7 -  val is_type_surely_infinite : Proof.context -> typ list -> typ -> bool
     1.8 +  val is_type_surely_infinite : Proof.context -> typ -> bool
     1.9    val monomorphic_term : Type.tyenv -> term -> term
    1.10    val eta_expand : typ list -> term -> int -> term
    1.11    val transform_elim_prop : term -> term
    1.12 @@ -136,70 +136,64 @@
    1.13     0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means
    1.14     cardinality 2 or more. The specified default cardinality is returned if the
    1.15     cardinality of the type can't be determined. *)
    1.16 -fun tiny_card_of_type ctxt default_card assigns T =
    1.17 +fun tiny_card_of_type ctxt default_card T =
    1.18    let
    1.19      val thy = Proof_Context.theory_of ctxt
    1.20      val max = 2 (* 1 would be too small for the "fun" case *)
    1.21      fun aux slack avoid T =
    1.22        if member (op =) avoid T then
    1.23          0
    1.24 -      else case AList.lookup (Sign.typ_instance thy o swap) assigns T of
    1.25 -        SOME k => k
    1.26 -      | NONE =>
    1.27 -        case T of
    1.28 -          Type (@{type_name fun}, [T1, T2]) =>
    1.29 -          (case (aux slack avoid T1, aux slack avoid T2) of
    1.30 -             (k, 1) => if slack andalso k = 0 then 0 else 1
    1.31 -           | (0, _) => 0
    1.32 -           | (_, 0) => 0
    1.33 -           | (k1, k2) =>
    1.34 -             if k1 >= max orelse k2 >= max then max
    1.35 -             else Int.min (max, Integer.pow k2 k1))
    1.36 -        | @{typ prop} => 2
    1.37 -        | @{typ bool} => 2 (* optimization *)
    1.38 -        | @{typ nat} => 0 (* optimization *)
    1.39 -        | Type ("Int.int", []) => 0 (* optimization *)
    1.40 -        | Type (s, _) =>
    1.41 -          (case datatype_constrs thy T of
    1.42 -             constrs as _ :: _ =>
    1.43 -             let
    1.44 -               val constr_cards =
    1.45 -                 map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
    1.46 -                      o snd) constrs
    1.47 -             in
    1.48 -               if exists (curry (op =) 0) constr_cards then 0
    1.49 -               else Int.min (max, Integer.sum constr_cards)
    1.50 -             end
    1.51 -           | [] =>
    1.52 -             case Typedef.get_info ctxt s of
    1.53 -               ({abs_type, rep_type, ...}, _) :: _ =>
    1.54 -               (* We cheat here by assuming that typedef types are infinite if
    1.55 -                  their underlying type is infinite. This is unsound in general
    1.56 -                  but it's hard to think of a realistic example where this would
    1.57 -                  not be the case. We are also slack with representation types:
    1.58 -                  If a representation type has the form "sigma => tau", we
    1.59 -                  consider it enough to check "sigma" for infiniteness. (Look
    1.60 -                  for "slack" in this function.) *)
    1.61 -               (case varify_and_instantiate_type ctxt
    1.62 -                         (Logic.varifyT_global abs_type) T
    1.63 -                         (Logic.varifyT_global rep_type)
    1.64 -                     |> aux true avoid of
    1.65 -                  0 => 0
    1.66 -                | 1 => 1
    1.67 -                | _ => default_card)
    1.68 -             | [] => default_card)
    1.69 -          (* Very slightly unsound: Type variables are assumed not to be
    1.70 -             constrained to cardinality 1. (In practice, the user would most
    1.71 -             likely have used "unit" directly anyway.) *)
    1.72 -        | TFree _ => if default_card = 1 then 2 else default_card
    1.73 -          (* Schematic type variables that contain only unproblematic sorts
    1.74 -             (with no finiteness axiom) can safely be considered infinite. *)
    1.75 -        | TVar _ => default_card
    1.76 +      else case T of
    1.77 +        Type (@{type_name fun}, [T1, T2]) =>
    1.78 +        (case (aux slack avoid T1, aux slack avoid T2) of
    1.79 +           (k, 1) => if slack andalso k = 0 then 0 else 1
    1.80 +         | (0, _) => 0
    1.81 +         | (_, 0) => 0
    1.82 +         | (k1, k2) =>
    1.83 +           if k1 >= max orelse k2 >= max then max
    1.84 +           else Int.min (max, Integer.pow k2 k1))
    1.85 +      | @{typ prop} => 2
    1.86 +      | @{typ bool} => 2 (* optimization *)
    1.87 +      | @{typ nat} => 0 (* optimization *)
    1.88 +      | Type ("Int.int", []) => 0 (* optimization *)
    1.89 +      | Type (s, _) =>
    1.90 +        (case datatype_constrs thy T of
    1.91 +           constrs as _ :: _ =>
    1.92 +           let
    1.93 +             val constr_cards =
    1.94 +               map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
    1.95 +                    o snd) constrs
    1.96 +           in
    1.97 +             if exists (curry (op =) 0) constr_cards then 0
    1.98 +             else Int.min (max, Integer.sum constr_cards)
    1.99 +           end
   1.100 +         | [] =>
   1.101 +           case Typedef.get_info ctxt s of
   1.102 +             ({abs_type, rep_type, ...}, _) :: _ =>
   1.103 +             (* We cheat here by assuming that typedef types are infinite if
   1.104 +                their underlying type is infinite. This is unsound in general
   1.105 +                but it's hard to think of a realistic example where this would
   1.106 +                not be the case. We are also slack with representation types:
   1.107 +                If a representation type has the form "sigma => tau", we
   1.108 +                consider it enough to check "sigma" for infiniteness. (Look
   1.109 +                for "slack" in this function.) *)
   1.110 +             (case varify_and_instantiate_type ctxt
   1.111 +                       (Logic.varifyT_global abs_type) T
   1.112 +                       (Logic.varifyT_global rep_type)
   1.113 +                   |> aux true avoid of
   1.114 +                0 => 0
   1.115 +              | 1 => 1
   1.116 +              | _ => default_card)
   1.117 +           | [] => default_card)
   1.118 +        (* Very slightly unsound: Type variables are assumed not to be
   1.119 +           constrained to cardinality 1. (In practice, the user would most
   1.120 +           likely have used "unit" directly anyway.) *)
   1.121 +      | TFree _ => if default_card = 1 then 2 else default_card
   1.122 +      | TVar _ => default_card
   1.123    in Int.min (max, aux false [] T) end
   1.124  
   1.125 -fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt 0 [] T <> 0
   1.126 -fun is_type_surely_infinite ctxt infinite_Ts T =
   1.127 -  tiny_card_of_type ctxt 1 (map (rpair 0) infinite_Ts) T = 0
   1.128 +fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt 0 T <> 0
   1.129 +fun is_type_surely_infinite ctxt T = tiny_card_of_type ctxt 1 T = 0
   1.130  
   1.131  fun monomorphic_term subst =
   1.132    map_types (map_type_tvar (fn v =>