|
1 (* Title: HOL/prod |
|
2 ID: $Id$ |
|
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 The type definition admits the following unused axiom: |
|
10 Abs_Unit_inverse "f: Unit ==> Rep_Unit(Abs_Unit(f)) = f" |
|
11 *) |
|
12 |
|
13 Prod = Set + |
|
14 types |
|
15 "*" 2 (infixr 20) |
|
16 unit 0 |
|
17 arities |
|
18 "*" :: (term,term)term |
|
19 unit :: term |
|
20 consts |
|
21 Pair_Rep :: "['a,'b] => ['a,'b] => bool" |
|
22 Prod :: "('a => 'b => bool)set" |
|
23 Rep_Prod :: "'a * 'b => ('a => 'b => bool)" |
|
24 Abs_Prod :: "('a => 'b => bool) => 'a * 'b" |
|
25 fst :: "'a * 'b => 'a" |
|
26 snd :: "'a * 'b => 'b" |
|
27 split :: "['a * 'b, ['a,'b]=>'c] => 'c" |
|
28 prod_fun :: "['a=>'b, 'c=>'d, 'a*'c] => 'b*'d" |
|
29 Pair :: "['a,'b] => 'a * 'b" |
|
30 "@Tuple" :: "args => 'a*'b" ("(1<_>)") |
|
31 Sigma :: "['a set, 'a => 'b set] => ('a*'b)set" |
|
32 |
|
33 Unit :: "bool set" |
|
34 Rep_Unit :: "unit => bool" |
|
35 Abs_Unit :: "bool => unit" |
|
36 Unity :: "unit" ("<>") |
|
37 |
|
38 translations |
|
39 |
|
40 "<x,y,z>" == "<x,<y,z>>" |
|
41 "<x,y>" == "Pair(x,y)" |
|
42 "<x>" => "x" |
|
43 |
|
44 rules |
|
45 |
|
46 Pair_Rep_def "Pair_Rep == (%a b. %x y. x=a & y=b)" |
|
47 Prod_def "Prod == {f. ? a b. f = Pair_Rep(a,b)}" |
|
48 (*faking a type definition...*) |
|
49 Rep_Prod "Rep_Prod(p): Prod" |
|
50 Rep_Prod_inverse "Abs_Prod(Rep_Prod(p)) = p" |
|
51 Abs_Prod_inverse "f: Prod ==> Rep_Prod(Abs_Prod(f)) = f" |
|
52 (*defining the abstract constants*) |
|
53 Pair_def "Pair(a,b) == Abs_Prod(Pair_Rep(a,b))" |
|
54 fst_def "fst(p) == @a. ? b. p = <a,b>" |
|
55 snd_def "snd(p) == @b. ? a. p = <a,b>" |
|
56 split_def "split(p,c) == c(fst(p),snd(p))" |
|
57 prod_fun_def "prod_fun(f,g) == (%z.split(z, %x y.<f(x), g(y)>))" |
|
58 Sigma_def "Sigma(A,B) == UN x:A. UN y:B(x). {<x,y>}" |
|
59 |
|
60 Unit_def "Unit == {p. p=True}" |
|
61 (*faking a type definition...*) |
|
62 Rep_Unit "Rep_Unit(u): Unit" |
|
63 Rep_Unit_inverse "Abs_Unit(Rep_Unit(u)) = u" |
|
64 (*defining the abstract constants*) |
|
65 Unity_def "Unity == Abs_Unit(True)" |
|
66 end |