| author | wenzelm | 
| Sat, 19 Aug 2000 12:44:39 +0200 | |
| changeset 9659 | b9cf6801f3da | 
| parent 9360 | 82e8b18e6985 | 
| permissions | -rw-r--r-- | 
| 923 | 1 | (* Title: HOL/Prod.thy | 
| 4570 | 2 | ID: $Id$ | 
| 923 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 4 | Copyright 1992 University of Cambridge | |
| 5 | ||
| 6 | Ordered Pairs and the Cartesian product type. | |
| 7 | The unit type. | |
| 8 | *) | |
| 9 | ||
| 1755 
17001ecd546e
Added additional parent theory equalities because some proofs in
 berghofe parents: 
1674diff
changeset | 10 | Prod = Fun + equalities + | 
| 923 | 11 | |
| 2260 | 12 | |
| 13 | (** products **) | |
| 923 | 14 | |
| 15 | (* type definition *) | |
| 16 | ||
| 1558 | 17 | constdefs | 
| 1370 
7361ac9b024d
removed quotes from types in consts and syntax sections
 clasohm parents: 
1273diff
changeset | 18 | Pair_Rep :: ['a, 'b] => ['a, 'b] => bool | 
| 1558 | 19 | "Pair_Rep == (%a b. %x y. x=a & y=b)" | 
| 923 | 20 | |
| 3947 | 21 | global | 
| 22 | ||
| 1475 | 23 | typedef (Prod) | 
| 923 | 24 |   ('a, 'b) "*"          (infixr 20)
 | 
| 25 |     = "{f. ? a b. f = Pair_Rep (a::'a) (b::'b)}"
 | |
| 26 | ||
| 2260 | 27 | syntax (symbols) | 
| 28 |   "*"           :: [type, type] => type         ("(_ \\<times>/ _)" [21, 20] 20)
 | |
| 29 | ||
| 6340 | 30 | syntax (HTML output) | 
| 31 |   "*"           :: [type, type] => type         ("(_ \\<times>/ _)" [21, 20] 20)
 | |
| 32 | ||
| 923 | 33 | |
| 34 | (* abstract constants and syntax *) | |
| 35 | ||
| 36 | consts | |
| 37 | fst :: "'a * 'b => 'a" | |
| 38 | snd :: "'a * 'b => 'b" | |
| 39 | split :: "[['a, 'b] => 'c, 'a * 'b] => 'c" | |
| 40 | prod_fun :: "['a => 'b, 'c => 'd, 'a * 'c] => 'b * 'd" | |
| 41 | Pair :: "['a, 'b] => 'a * 'b" | |
| 42 |   Sigma         :: "['a set, 'a => 'b set] => ('a * 'b) set"
 | |
| 43 | ||
| 2260 | 44 | |
| 45 | (* patterns -- extends pre-defined type "pttrn" used in abstractions *) | |
| 46 | ||
| 4875 | 47 | nonterminals | 
| 9360 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 48 | tuple_args patterns | 
| 1068 | 49 | |
| 923 | 50 | syntax | 
| 9360 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 51 |   "_tuple"      :: "'a => tuple_args => 'a * 'b"        ("(1'(_,/ _'))")
 | 
| 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 52 |   "_tuple_arg"  :: "'a => tuple_args"                   ("_")
 | 
| 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 53 |   "_tuple_args" :: "'a => tuple_args => tuple_args"     ("_,/ _")
 | 
| 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 54 |   "_pattern"    :: [pttrn, patterns] => pttrn           ("'(_,/ _')")
 | 
| 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 55 |   ""            :: pttrn => patterns                    ("_")
 | 
| 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 56 |   "_patterns"   :: [pttrn, patterns] => patterns        ("_,/ _")
 | 
| 3692 | 57 |   "@Sigma"      :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"   ("(3SIGMA _:_./ _)" 10)
 | 
| 8703 | 58 |   "@Times"      :: "['a set, 'a => 'b set] => ('a * 'b) set"    (infixr "<*>" 80)
 | 
| 1636 | 59 | |
| 923 | 60 | translations | 
| 3842 | 61 | "(x, y)" == "Pair x y" | 
| 9360 
82e8b18e6985
more robust tuple syntax (still improper, though!);
 wenzelm parents: 
9341diff
changeset | 62 | "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))" | 
| 3842 | 63 | "%(x,y,zs).b" == "split(%x (y,zs).b)" | 
| 64 | "%(x,y).b" == "split(%x y. b)" | |
| 2973 
184c7cd8043d
Added ability to have case expressions involving tuples. (via translation)
 nipkow parents: 
2886diff
changeset | 65 | "_abs (Pair x y) t" => "%(x,y).t" | 
| 
184c7cd8043d
Added ability to have case expressions involving tuples. (via translation)
 nipkow parents: 
2886diff
changeset | 66 | (* The last rule accommodates tuples in `case C ... (x,y) ... => ...' | 
| 
184c7cd8043d
Added ability to have case expressions involving tuples. (via translation)
 nipkow parents: 
2886diff
changeset | 67 | The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *) | 
| 2260 | 68 | |
| 3842 | 69 | "SIGMA x:A. B" => "Sigma A (%x. B)" | 
| 8703 | 70 | "A <*> B" => "Sigma A (_K B)" | 
| 1068 | 71 | |
| 2260 | 72 | syntax (symbols) | 
| 3692 | 73 |   "@Sigma"      :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"   ("(3\\<Sigma> _\\<in>_./ _)" 10)
 | 
| 2260 | 74 |   "@Times"      :: "['a set, 'a => 'b set] => ('a * 'b) set"    ("_ \\<times> _" [81, 80] 80)
 | 
| 75 | ||
| 76 | ||
| 77 | (* definitions *) | |
| 1636 | 78 | |
| 3947 | 79 | local | 
| 80 | ||
| 923 | 81 | defs | 
| 82 | Pair_def "Pair a b == Abs_Prod(Pair_Rep a b)" | |
| 2393 | 83 | fst_def "fst p == @a. ? b. p = (a, b)" | 
| 84 | snd_def "snd p == @b. ? a. p = (a, b)" | |
| 1655 | 85 | split_def "split == (%c p. c (fst p) (snd p))" | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
967diff
changeset | 86 | prod_fun_def "prod_fun f g == split(%x y.(f(x), g(y)))" | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
967diff
changeset | 87 |   Sigma_def     "Sigma A B == UN x:A. UN y:B(x). {(x, y)}"
 | 
| 923 | 88 | |
| 2260 | 89 | |
| 90 | ||
| 91 | (** unit **) | |
| 923 | 92 | |
| 3947 | 93 | global | 
| 94 | ||
| 2886 | 95 | typedef  unit = "{True}"
 | 
| 923 | 96 | |
| 97 | consts | |
| 1370 
7361ac9b024d
removed quotes from types in consts and syntax sections
 clasohm parents: 
1273diff
changeset | 98 |   "()"          :: unit                           ("'(')")
 | 
| 923 | 99 | |
| 3947 | 100 | local | 
| 101 | ||
| 923 | 102 | defs | 
| 2880 | 103 | Unity_def "() == Abs_unit True" | 
| 1273 | 104 | |
| 923 | 105 | end | 
| 1636 | 106 | |
| 107 | ML | |
| 108 | ||
| 109 | val print_translation = [("Sigma", dependent_tr' ("@Sigma", "@Times"))];
 |