| author | wenzelm | 
| Mon, 12 Jul 1999 22:25:39 +0200 | |
| changeset 6980 | bb526ba7ba5f | 
| parent 6340 | 7d5cbd5819a0 | 
| child 8703 | 816d8f6513be | 
| 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: 
1674 
diff
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: 
1273 
diff
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  | 
48  | 
patterns  | 
|
| 1068 | 49  | 
|
| 923 | 50  | 
syntax  | 
| 3692 | 51  | 
  "@Tuple"      :: "['a, args] => 'a * 'b"       ("(1'(_,/ _'))")
 | 
| 923 | 52  | 
|
| 3692 | 53  | 
  "_pattern"    :: [pttrn, patterns] => pttrn    ("'(_,/_')")
 | 
54  | 
  ""            :: pttrn => patterns             ("_")
 | 
|
55  | 
  "_patterns"   :: [pttrn, patterns] => patterns ("_,/_")
 | 
|
| 
2973
 
184c7cd8043d
Added ability to have case expressions involving tuples. (via translation)
 
nipkow 
parents: 
2886 
diff
changeset
 | 
56  | 
|
| 3692 | 57  | 
  "@Sigma"      :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"   ("(3SIGMA _:_./ _)" 10)
 | 
| 2260 | 58  | 
  "@Times"      :: "['a set, 'a => 'b set] => ('a * 'b) set"    ("_ Times _" [81, 80] 80)
 | 
| 1636 | 59  | 
|
| 923 | 60  | 
translations  | 
| 3842 | 61  | 
"(x, y, z)" == "(x, (y, z))"  | 
62  | 
"(x, y)" == "Pair x y"  | 
|
| 923 | 63  | 
|
| 3842 | 64  | 
"%(x,y,zs).b" == "split(%x (y,zs).b)"  | 
65  | 
"%(x,y).b" == "split(%x y. b)"  | 
|
| 
2973
 
184c7cd8043d
Added ability to have case expressions involving tuples. (via translation)
 
nipkow 
parents: 
2886 
diff
changeset
 | 
66  | 
"_abs (Pair x y) t" => "%(x,y).t"  | 
| 
 
184c7cd8043d
Added ability to have case expressions involving tuples. (via translation)
 
nipkow 
parents: 
2886 
diff
changeset
 | 
67  | 
(* The last rule accommodates tuples in `case C ... (x,y) ... => ...'  | 
| 
 
184c7cd8043d
Added ability to have case expressions involving tuples. (via translation)
 
nipkow 
parents: 
2886 
diff
changeset
 | 
68  | 
The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *)  | 
| 2260 | 69  | 
|
| 3842 | 70  | 
"SIGMA x:A. B" => "Sigma A (%x. B)"  | 
71  | 
"A Times B" => "Sigma A (_K B)"  | 
|
| 1068 | 72  | 
|
| 2260 | 73  | 
syntax (symbols)  | 
| 3692 | 74  | 
  "@Sigma"      :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"   ("(3\\<Sigma> _\\<in>_./ _)" 10)
 | 
| 2260 | 75  | 
  "@Times"      :: "['a set, 'a => 'b set] => ('a * 'b) set"    ("_ \\<times> _" [81, 80] 80)
 | 
76  | 
||
77  | 
||
78  | 
(* definitions *)  | 
|
| 1636 | 79  | 
|
| 3947 | 80  | 
local  | 
81  | 
||
| 923 | 82  | 
defs  | 
83  | 
Pair_def "Pair a b == Abs_Prod(Pair_Rep a b)"  | 
|
| 2393 | 84  | 
fst_def "fst p == @a. ? b. p = (a, b)"  | 
85  | 
snd_def "snd p == @b. ? a. p = (a, b)"  | 
|
| 1655 | 86  | 
split_def "split == (%c p. c (fst p) (snd p))"  | 
| 
972
 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 
clasohm 
parents: 
967 
diff
changeset
 | 
87  | 
prod_fun_def "prod_fun f g == split(%x y.(f(x), g(y)))"  | 
| 
 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 
clasohm 
parents: 
967 
diff
changeset
 | 
88  | 
  Sigma_def     "Sigma A B == UN x:A. UN y:B(x). {(x, y)}"
 | 
| 923 | 89  | 
|
| 2260 | 90  | 
|
91  | 
||
92  | 
(** unit **)  | 
|
| 923 | 93  | 
|
| 3947 | 94  | 
global  | 
95  | 
||
| 2886 | 96  | 
typedef  unit = "{True}"
 | 
| 923 | 97  | 
|
98  | 
consts  | 
|
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
99  | 
  "()"          :: unit                           ("'(')")
 | 
| 923 | 100  | 
|
| 3947 | 101  | 
local  | 
102  | 
||
| 923 | 103  | 
defs  | 
| 2880 | 104  | 
Unity_def "() == Abs_unit True"  | 
| 1273 | 105  | 
|
| 923 | 106  | 
end  | 
| 1636 | 107  | 
|
108  | 
ML  | 
|
109  | 
||
110  | 
val print_translation = [("Sigma", dependent_tr' ("@Sigma", "@Times"))];
 |