src/HOL/Prod.ML
author clasohm
Wed, 04 Oct 1995 13:10:03 +0100
changeset 1264 3eb91524b938
parent 972 e61b058d58d2
child 1301 42782316d510
permissions -rw-r--r--
added local simpsets; removed IOA from 'make test'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/prod
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
For prod.thy.  Ordered Pairs, the Cartesian product type, the unit type
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
open Prod;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
(*This counts as a non-emptiness result for admitting 'a * 'b as a type*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
goalw Prod.thy [Prod_def] "Pair_Rep a b : Prod";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
by (EVERY1 [rtac CollectI, rtac exI, rtac exI, rtac refl]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
qed "ProdI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
val [major] = goalw Prod.thy [Pair_Rep_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
    "Pair_Rep a b = Pair_Rep a' b' ==> a=a' & b=b'";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
by (EVERY1 [rtac (major RS fun_cong RS fun_cong RS subst), 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
	    rtac conjI, rtac refl, rtac refl]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
qed "Pair_Rep_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
goal Prod.thy "inj_onto Abs_Prod Prod";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
by (rtac inj_onto_inverseI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
by (etac Abs_Prod_inverse 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
qed "inj_onto_Abs_Prod";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
val prems = goalw Prod.thy [Pair_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    28
    "[| (a, b) = (a',b');  [| a=a';  b=b' |] ==> R |] ==> R";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
by (rtac (inj_onto_Abs_Prod RS inj_ontoD RS Pair_Rep_inject RS conjE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
by (REPEAT (ares_tac (prems@[ProdI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
qed "Pair_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    33
goal Prod.thy "((a,b) = (a',b')) = (a=a' & b=b')";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
by (fast_tac (set_cs addIs [Pair_inject]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
qed "Pair_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    37
goalw Prod.thy [fst_def] "fst((a,b)) = a";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
qed "fst_conv";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    41
goalw Prod.thy [snd_def] "snd((a,b)) = b";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
qed "snd_conv";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    45
goalw Prod.thy [Pair_def] "? x y. p = (x,y)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
by (EVERY1[etac exE, etac exE, rtac exI, rtac exI,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
	   rtac (Rep_Prod_inverse RS sym RS trans),  etac arg_cong]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
qed "PairE_lemma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    51
val [prem] = goal Prod.thy "[| !!x y. p = (x,y) ==> Q |] ==> Q";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
by (rtac (PairE_lemma RS exE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
by (REPEAT (eresolve_tac [prem,exE] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
qed "PairE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    56
goalw Prod.thy [split_def] "split c (a,b) = c a b";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
by (sstac [fst_conv, snd_conv] 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
by (rtac refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
qed "split";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
    61
Addsimps [fst_conv, snd_conv, split, Pair_eq];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
goal Prod.thy "(s=t) = (fst(s)=fst(t) & snd(s)=snd(t))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
by (res_inst_tac[("p","s")] PairE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
by (res_inst_tac[("p","t")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
    66
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
qed "Pair_fst_snd_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
(*Prevents simplification of c: much faster*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
qed_goal "split_weak_cong" Prod.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
  "p=q ==> split c p = split c q"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
  (fn [prem] => [rtac (prem RS arg_cong) 1]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
(* Do not add as rewrite rule: invalidates some proofs in IMP *)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    75
goal Prod.thy "p = (fst(p),snd(p))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
by (res_inst_tac [("p","p")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
    77
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
qed "surjective_pairing";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    80
goal Prod.thy "p = split (%x y.(x,y)) p";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
by (res_inst_tac [("p","p")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
    82
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
qed "surjective_pairing2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
(*For use with split_tac and the simplifier*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    86
goal Prod.thy "R(split c p) = (! x y. p = (x,y) --> R(c x y))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
by (stac surjective_pairing 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
by (stac split 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
by (fast_tac (HOL_cs addSEs [Pair_inject]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
qed "expand_split";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
(** split used as a logical connective or set former **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
(*These rules are for use with fast_tac.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
  Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    97
goal Prod.thy "!!a b c. c a b ==> split c (a,b)";
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
    98
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
qed "splitI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
val prems = goalw Prod.thy [split_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   102
    "[| split c p;  !!x y. [| p = (x,y);  c x y |] ==> Q |] ==> Q";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
qed "splitE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   106
goal Prod.thy "!!R a b. split R (a,b) ==> R a b";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
by (etac (split RS iffD1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
qed "splitD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   110
goal Prod.thy "!!a b c. z: c a b ==> z: split c (a,b)";
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   111
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
qed "mem_splitI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
val prems = goalw Prod.thy [split_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   115
    "[| z: split c p;  !!x y. [| p = (x,y);  z: c x y |] ==> Q |] ==> Q";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
qed "mem_splitE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
(*** prod_fun -- action of the product functor upon functions ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   121
goalw Prod.thy [prod_fun_def] "prod_fun f g (a,b) = (f(a),g(b))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
by (rtac split 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
qed "prod_fun";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
goal Prod.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
    "prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
by (rtac ext 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
by (res_inst_tac [("p","x")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   129
by (asm_simp_tac (!simpset addsimps [prod_fun,o_def]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
qed "prod_fun_compose";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
goal Prod.thy "prod_fun (%x.x) (%y.y) = (%z.z)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
by (rtac ext 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
by (res_inst_tac [("p","z")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   135
by (asm_simp_tac (!simpset addsimps [prod_fun]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
qed "prod_fun_ident";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   138
val prems = goal Prod.thy "(a,b):r ==> (f(a),g(b)) : (prod_fun f g)``r";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
by (rtac image_eqI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
by (rtac (prod_fun RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
qed "prod_fun_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
val major::prems = goal Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   145
    "[| c: (prod_fun f g)``r;  !!x y. [| c=(f(x),g(y));  (x,y):r |] ==> P  \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   146
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
by (rtac (major RS imageE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   148
by (res_inst_tac [("p","x")] PairE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
by (fast_tac HOL_cs 2);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
by (fast_tac (HOL_cs addIs [prod_fun]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
qed "prod_fun_imageE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
(*** Disjoint union of a family of sets - Sigma ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
qed_goalw "SigmaI" Prod.thy [Sigma_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   157
    "[| a:A;  b:B(a) |] ==> (a,b) : Sigma A B"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   158
 (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   160
(*The general elimination rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
qed_goalw "SigmaE" Prod.thy [Sigma_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
    "[| c: Sigma A B;  \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   163
\       !!x y.[| x:A;  y:B(x);  c=(x,y) |] ==> P \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
\    |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
 (fn major::prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
  [ (cut_facts_tac [major] 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
    (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   169
(** Elimination of (a,b):A*B -- introduces no eigenvariables **)
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   170
qed_goal "SigmaD1" Prod.thy "(a,b) : Sigma A B ==> a : A"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   171
 (fn [major]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   172
  [ (rtac (major RS SigmaE) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   173
    (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   174
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   175
qed_goal "SigmaD2" Prod.thy "(a,b) : Sigma A B ==> b : B(a)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
 (fn [major]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
  [ (rtac (major RS SigmaE) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
    (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   179
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
qed_goal "SigmaE2" Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   181
    "[| (a,b) : Sigma A B;    \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
\       [| a:A;  b:B(a) |] ==> P   \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
\    |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
 (fn [major,minor]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
  [ (rtac minor 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
    (rtac (major RS SigmaD1) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
    (rtac (major RS SigmaD2) 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
(*** Domain of a relation ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   191
val prems = goalw Prod.thy [image_def] "(a,b) : r ==> a : fst``r";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
by (rtac CollectI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
by (rtac bexI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
by (rtac (fst_conv RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
qed "fst_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
val major::prems = goal Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   199
    "[| a : fst``r;  !!y.[| (a,y) : r |] ==> P |] ==> P"; 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
by (rtac (major RS imageE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
by (etac ssubst 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   203
by (rtac (surjective_pairing RS subst) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
by (assume_tac 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
qed "fst_imageE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
(*** Range of a relation ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   208
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   209
val prems = goalw Prod.thy [image_def] "(a,b) : r ==> b : snd``r";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
by (rtac CollectI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
by (rtac bexI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   212
by (rtac (snd_conv RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   213
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   214
qed "snd_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
val major::prems = goal Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   217
    "[| a : snd``r;  !!y.[| (y,a) : r |] ==> P |] ==> P"; 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   218
by (rtac (major RS imageE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   220
by (etac ssubst 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   221
by (rtac (surjective_pairing RS subst) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   222
by (assume_tac 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   223
qed "snd_imageE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   224
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   225
(** Exhaustion rule for unit -- a degenerate form of induction **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   226
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   227
goalw Prod.thy [Unity_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   228
    "u = ()";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   229
by (stac (rewrite_rule [Unit_def] Rep_Unit RS CollectD RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   230
by (rtac (Rep_Unit_inverse RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
qed "unit_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   232
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   233
val prod_cs = set_cs addSIs [SigmaI, mem_splitI] 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   234
                     addIs  [fst_imageI, snd_imageI, prod_fun_imageI]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
                     addSEs [SigmaE2, SigmaE, mem_splitE, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   236
			     fst_imageE, snd_imageE, prod_fun_imageE,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   237
			     Pair_inject];