src/ZF/cartprod.ML
author wenzelm
Sun, 29 Nov 1998 13:16:47 +0100
changeset 5989 9670dae0143d
parent 3397 3e2b8d0de2a0
permissions -rw-r--r--
proof_general_trans (experimental);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     1
(*  Title:      ZF/cartprod.ML
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     2
    ID:         $Id$
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     5
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     6
Syntactic operations for Cartesian Products
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     7
*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     8
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
     9
signature PR =                  (** Description of a Cartesian product **)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    10
  sig
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    11
  val sigma     : term                  (*Cartesian product operator*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    12
  val pair      : term                  (*pairing operator*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    13
  val split_name : string               (*name of polymorphic split*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    14
  val pair_iff  : thm                   (*injectivity of pairing, using <->*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    15
  val split_eq  : thm                   (*equality rule for split*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    16
  val fsplitI   : thm                   (*intro rule for fsplit*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    17
  val fsplitD   : thm                   (*destruct rule for fsplit*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    18
  val fsplitE   : thm                   (*elim rule; apparently never used*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    19
  end;
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    20
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    21
signature CARTPROD =            (** Derived syntactic functions for produts **)
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    22
  sig
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    23
  val ap_split : typ -> typ -> term -> term
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    24
  val factors : typ -> typ list
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    25
  val mk_prod : typ * typ -> typ
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    26
  val mk_tuple : term -> typ -> term list -> term
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    27
  val pseudo_type : term -> typ
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    28
  val remove_split : thm -> thm
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    29
  val split_const : typ -> term
3397
3e2b8d0de2a0 Made the pseudo-type of split_rule_var a separate argument
paulson
parents: 2033
diff changeset
    30
  val split_rule_var : term * typ * thm -> thm
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    31
  end;
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    32
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    33
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    34
functor CartProd_Fun (Pr: PR) : CARTPROD =
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    35
struct
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    36
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    37
(* Some of these functions expect "pseudo-types" containing products,
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    38
   as in HOL; the true ZF types would just be "i" *)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    39
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    40
fun mk_prod (T1,T2) = Type("*", [T1,T2]);
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    41
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    42
(*Bogus product type underlying a (possibly nested) Sigma.  
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    43
  Lets us share HOL code*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    44
fun pseudo_type (t $ A $ Abs(_,_,B)) = 
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    45
        if t = Pr.sigma  then  mk_prod(pseudo_type A, pseudo_type B)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    46
                         else  Ind_Syntax.iT
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    47
  | pseudo_type _           =  Ind_Syntax.iT;
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    48
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    49
(*Maps the type T1*...*Tn to [T1,...,Tn], however nested*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    50
fun factors (Type("*", [T1,T2])) = factors T1 @ factors T2
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    51
  | factors T                    = [T];
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    52
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    53
(*Make a well-typed instance of "split"*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    54
fun split_const T = Const(Pr.split_name, 
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    55
                          [[Ind_Syntax.iT, Ind_Syntax.iT]--->T, 
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    56
                           Ind_Syntax.iT] ---> T);
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    57
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    58
(*In ap_split S T u, term u expects separate arguments for the factors of S,
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    59
  with result type T.  The call creates a new term expecting one argument
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    60
  of type S.*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    61
fun ap_split (Type("*", [T1,T2])) T3 u   = 
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    62
       split_const T3 $ 
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    63
       Abs("v", Ind_Syntax.iT, (*Not T1, as it involves pseudo-product types*)
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    64
           ap_split T2 T3
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    65
           ((ap_split T1 (factors T2 ---> T3) (incr_boundvars 1 u)) $ 
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    66
            Bound 0))
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    67
  | ap_split T T3 u = u;
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    68
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    69
(*Makes a nested tuple from a list, following the product type structure*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    70
fun mk_tuple pair (Type("*", [T1,T2])) tms = 
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    71
        pair $ (mk_tuple pair T1 tms)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    72
             $ (mk_tuple pair T2 (drop (length (factors T1), tms)))
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    73
  | mk_tuple pair T (t::_) = t;
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    74
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    75
(*Attempts to remove occurrences of split, and pair-valued parameters*)
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    76
val remove_split = rewrite_rule [Pr.split_eq];
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    77
3397
3e2b8d0de2a0 Made the pseudo-type of split_rule_var a separate argument
paulson
parents: 2033
diff changeset
    78
(*Uncurries any Var according to its "pseudo-product type" T1 in the rule*)
3e2b8d0de2a0 Made the pseudo-type of split_rule_var a separate argument
paulson
parents: 2033
diff changeset
    79
fun split_rule_var (Var(v,_), Type("fun",[T1,T2]), rl) =
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    80
      let val T' = factors T1 ---> T2
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    81
          val newt = ap_split T1 T2 (Var(v,T'))
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    82
          val cterm = Thm.cterm_of (#sign(rep_thm rl))
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    83
      in
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    84
          remove_split (instantiate ([], [(cterm (Var(v, Ind_Syntax.iT-->T2)), 
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1734
diff changeset
    85
                                           cterm newt)]) rl)
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    86
      end
3397
3e2b8d0de2a0 Made the pseudo-type of split_rule_var a separate argument
paulson
parents: 2033
diff changeset
    87
  | split_rule_var (t,T,rl) = rl;
1734
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    88
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    89
end;
604da1a11a99 New functor for operating on the two forms of Cartesian product
paulson
parents:
diff changeset
    90