src/HOL/Prod.ML
author wenzelm
Mon, 01 Dec 1997 18:22:38 +0100
changeset 4334 e567f3425267
parent 4192 c38ab5af38b5
child 4423 a129b817b58a
permissions -rw-r--r--
ISABELLE_TMP_PREFIX;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
     1
(*  Title:      HOL/prod
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
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), 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
    19
            rtac conjI, rtac refl, rtac refl]);
923
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')";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    34
by (blast_tac (claset() addSEs [Pair_inject]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
qed "Pair_eq";
3429
nipkow
parents: 2935
diff changeset
    36
AddIffs [Pair_eq];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    38
goalw Prod.thy [fst_def] "fst((a,b)) = a";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    39
by (blast_tac (claset() addIs [select_equality]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
qed "fst_conv";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    42
goalw Prod.thy [snd_def] "snd((a,b)) = b";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    43
by (blast_tac (claset() addIs [select_equality]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
qed "snd_conv";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    46
goalw Prod.thy [Pair_def] "? x y. p = (x,y)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
by (EVERY1[etac exE, etac exE, rtac exI, rtac exI,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
    49
           rtac (Rep_Prod_inverse RS sym RS trans),  etac arg_cong]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
qed "PairE_lemma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    52
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
    53
by (rtac (PairE_lemma RS exE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
by (REPEAT (eresolve_tac [prem,exE] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
qed "PairE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
    57
fun pair_tac s = res_inst_tac [("p",s)] PairE THEN' hyp_subst_tac;
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
    58
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    59
(* replace parameters of product type by individual component parameters *)
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    60
local
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    61
fun is_pair (_,Type("*",_)) = true
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    62
  | is_pair _ = false;
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    63
1727
7d0fbdc46e8e Now split_all_tac works for i>1 !
paulson
parents: 1655
diff changeset
    64
fun find_pair_param prem =
7d0fbdc46e8e Now split_all_tac works for i>1 !
paulson
parents: 1655
diff changeset
    65
  let val params = Logic.strip_params prem
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    66
  in if exists is_pair params
1727
7d0fbdc46e8e Now split_all_tac works for i>1 !
paulson
parents: 1655
diff changeset
    67
     then let val params = rev(rename_wrt_term prem params)
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    68
                           (*as they are printed*)
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    69
          in apsome fst (find_first is_pair params) end
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    70
     else None
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    71
  end;
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    72
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    73
in
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    74
1727
7d0fbdc46e8e Now split_all_tac works for i>1 !
paulson
parents: 1655
diff changeset
    75
val split_all_tac = REPEAT o SUBGOAL (fn (prem,i) =>
7d0fbdc46e8e Now split_all_tac works for i>1 !
paulson
parents: 1655
diff changeset
    76
  case find_pair_param prem of
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    77
    None => no_tac
1727
7d0fbdc46e8e Now split_all_tac works for i>1 !
paulson
parents: 1655
diff changeset
    78
  | Some x => EVERY[res_inst_tac[("p",x)] PairE i,
7d0fbdc46e8e Now split_all_tac works for i>1 !
paulson
parents: 1655
diff changeset
    79
                    REPEAT(hyp_subst_tac i), prune_params_tac]);
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    80
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    81
end;
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
    82
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    83
(* Could be nice, but breaks too many proofs:
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    84
claset_ref() := claset() addbefore split_all_tac;
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    85
*)
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    86
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    87
(*** lemmas for splitting paired `!!'
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    88
Does not work with simplifier because it also affects premises in
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    89
congrence rules, where is can lead to premises of the form
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    90
!!a b. ... = ?P(a,b)
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    91
which cannot be solved by reflexivity.
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    92
   
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3568
diff changeset
    93
val [prem] = goal Prod.thy "(!!x. PROP P x) ==> (!!a b. PROP P(a,b))";
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    94
br prem 1;
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    95
val lemma1 = result();
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    96
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    97
local
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    98
    val psig = sign_of Prod.thy;
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
    99
    val pT = Sign.read_typ (psig, K None) "?'a*?'b=>prop";
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   100
    val PeqP = reflexive(read_cterm psig ("P", pT));
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   101
    val psplit = zero_var_indexes(read_instantiate [("p","x")]
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   102
                                  surjective_pairing RS eq_reflection)
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   103
in
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   104
val adhoc = combination PeqP psplit
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   105
end;
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   106
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   107
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   108
val [prem] = goal Prod.thy "(!!a b. PROP P(a,b)) ==> PROP P x";
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   109
bw adhoc;
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   110
br prem 1;
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   111
val lemma = result();
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   112
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   113
val [prem] = goal Prod.thy "(!!a b. PROP P(a,b)) ==> (!!x. PROP P x)";
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   114
br lemma 1;
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   115
br prem 1;
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   116
val lemma2 = result();
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   117
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   118
bind_thm("split_paired_all", equal_intr lemma1 lemma2);
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   119
Addsimps [split_paired_all];
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   120
***)
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   121
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   122
goal Prod.thy "(!x. P x) = (!a b. P(a,b))";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   123
by (fast_tac (claset() addbefore split_all_tac) 1);
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   124
qed "split_paired_All";
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   125
Addsimps [split_paired_All];
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   126
(* AddIffs is not a good idea because it makes Blast_tac loop *)
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   127
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   128
goal Prod.thy "(? x. P x) = (? a b. P(a,b))";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   129
by (fast_tac (claset() addbefore split_all_tac) 1);
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   130
qed "split_paired_Ex";
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   131
(* Addsimps [split_paired_Ex]; breaks a number of IOA proofs *)
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   132
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   133
goalw Prod.thy [split_def] "split c (a,b) = c a b";
1485
240cc98b94a7 Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents: 1465
diff changeset
   134
by (EVERY1[stac fst_conv, stac snd_conv]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
by (rtac refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
qed "split";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   138
val split_select = prove_goalw Prod.thy [split_def] 
4192
c38ab5af38b5 replaced 8bit characters
oheimb
parents: 4134
diff changeset
   139
	"(@(x,y). P x y) = (@xy. P (fst xy) (snd xy))" (K [rtac refl 1]);
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   140
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   141
Addsimps [fst_conv, snd_conv, split];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
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
   144
by (res_inst_tac[("p","s")] PairE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145
by (res_inst_tac[("p","t")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   146
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
qed "Pair_fst_snd_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   148
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
(*Prevents simplification of c: much faster*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
qed_goal "split_weak_cong" Prod.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
  "p=q ==> split c p = split c q"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
  (fn [prem] => [rtac (prem RS arg_cong) 1]);
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
(* Do not add as rewrite rule: invalidates some proofs in IMP *)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   155
goal Prod.thy "p = (fst(p),snd(p))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
by (res_inst_tac [("p","p")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   157
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   158
qed "surjective_pairing";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   160
goal Prod.thy "p = split (%x y.(x,y)) p";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
by (res_inst_tac [("p","p")] PairE 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   162
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
qed "surjective_pairing2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   165
val surj_pair = prove_goal Prod.thy "? x y. z = (x, y)" (K [
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   166
	rtac exI 1, rtac exI 1, rtac surjective_pairing 1]);
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   167
Addsimps [surj_pair];
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   168
1655
5be64540f275 Added a number of lemmas
nipkow
parents: 1642
diff changeset
   169
qed_goal "split_eta" Prod.thy "(%(x,y). f(x,y)) = f"
5be64540f275 Added a number of lemmas
nipkow
parents: 1642
diff changeset
   170
  (fn _ => [rtac ext 1, split_all_tac 1, rtac split 1]);
5be64540f275 Added a number of lemmas
nipkow
parents: 1642
diff changeset
   171
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   172
val split_beta = prove_goal Prod.thy "(%(x,y). P x y) z = P (fst z) (snd z)"
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   173
	(fn _ => [stac surjective_pairing 1, stac split 1, rtac refl 1]);
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   174
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
(*For use with split_tac and the simplifier*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   176
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
   177
by (stac surjective_pairing 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
by (stac split 1);
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   179
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
qed "expand_split";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   182
(* could be done after split_tac has been speeded up significantly:
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   183
simpset_ref() := (simpset() addsplits [expand_split]);
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   184
   precompute the constants involved and don't do anything unless
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   185
   the current goal contains one of those constants
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   186
*)
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   187
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
(** split used as a logical connective or set former **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   190
(*These rules are for use with blast_tac.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
  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
   192
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   193
goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> c a b |] ==> split c p";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1515
diff changeset
   194
by (split_all_tac 1);
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   195
by (Asm_simp_tac 1);
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   196
qed "splitI2";
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   197
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   198
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
   199
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
qed "splitI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
val prems = goalw Prod.thy [split_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   203
    "[| 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
   204
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
qed "splitE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   207
val splitE2 = prove_goal Prod.thy 
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   208
"[|Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R|] ==> R" (fn prems => [
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   209
	REPEAT (resolve_tac (prems@[surjective_pairing]) 1),
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   210
	rtac (split_beta RS subst) 1,
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   211
	rtac (hd prems) 1]);
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   212
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   213
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
   214
by (etac (split RS iffD1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
qed "splitD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   217
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
   218
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
qed "mem_splitI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   220
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   221
goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> z: c a b |] ==> z: split c p";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1515
diff changeset
   222
by (split_all_tac 1);
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   223
by (Asm_simp_tac 1);
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   224
qed "mem_splitI2";
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   225
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   226
val prems = goalw Prod.thy [split_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   227
    "[| 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
   228
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   229
qed "mem_splitE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   230
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   231
AddSIs [splitI, splitI2, mem_splitI, mem_splitI2];
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   232
AddSEs [splitE, mem_splitE];
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   233
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   234
(*** prod_fun -- action of the product functor upon functions ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   236
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
   237
by (rtac split 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   238
qed "prod_fun";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   239
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   240
goal Prod.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   241
    "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
   242
by (rtac ext 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   243
by (res_inst_tac [("p","x")] PairE 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   244
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
   245
qed "prod_fun_compose";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   246
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3568
diff changeset
   247
goal Prod.thy "prod_fun (%x. x) (%y. y) = (%z. z)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   248
by (rtac ext 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   249
by (res_inst_tac [("p","z")] PairE 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   250
by (asm_simp_tac (simpset() addsimps [prod_fun]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   251
qed "prod_fun_ident";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   252
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   253
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
   254
by (rtac image_eqI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   255
by (rtac (prod_fun RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   256
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   257
qed "prod_fun_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   258
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   259
val major::prems = goal Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   260
    "[| 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
   261
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   262
by (rtac (major RS imageE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   263
by (res_inst_tac [("p","x")] PairE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   264
by (resolve_tac prems 1);
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   265
by (Blast_tac 2);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   266
by (blast_tac (claset() addIs [prod_fun]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   267
qed "prod_fun_imageE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   268
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   269
(*** Disjoint union of a family of sets - Sigma ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   270
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   271
qed_goalw "SigmaI" Prod.thy [Sigma_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   272
    "[| a:A;  b:B(a) |] ==> (a,b) : Sigma A B"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   273
 (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   274
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   275
AddSIs [SigmaI];
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   276
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   277
(*The general elimination rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   278
qed_goalw "SigmaE" Prod.thy [Sigma_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   279
    "[| c: Sigma A B;  \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   280
\       !!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
   281
\    |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   282
 (fn major::prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   283
  [ (cut_facts_tac [major] 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   284
    (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
   285
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   286
(** Elimination of (a,b):A*B -- introduces no eigenvariables **)
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   287
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
   288
 (fn [major]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   289
  [ (rtac (major RS SigmaE) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   290
    (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   291
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   292
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
   293
 (fn [major]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   294
  [ (rtac (major RS SigmaE) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   295
    (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   296
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   297
qed_goal "SigmaE2" Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   298
    "[| (a,b) : Sigma A B;    \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   299
\       [| a:A;  b:B(a) |] ==> P   \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   300
\    |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   301
 (fn [major,minor]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   302
  [ (rtac minor 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   303
    (rtac (major RS SigmaD1) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   304
    (rtac (major RS SigmaD2) 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   305
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   306
AddSEs [SigmaE2, SigmaE];
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   307
1515
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   308
val prems = goal Prod.thy
1642
21db0cf9a1a4 Using new "Times" infix
paulson
parents: 1618
diff changeset
   309
    "[| A<=C;  !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D";
1515
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   310
by (cut_facts_tac prems 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   311
by (blast_tac (claset() addIs (prems RL [subsetD])) 1);
1515
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   312
qed "Sigma_mono";
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   313
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   314
qed_goal "Sigma_empty1" Prod.thy "Sigma {} B = {}"
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   315
 (fn _ => [ (Blast_tac 1) ]);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   316
1642
21db0cf9a1a4 Using new "Times" infix
paulson
parents: 1618
diff changeset
   317
qed_goal "Sigma_empty2" Prod.thy "A Times {} = {}"
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   318
 (fn _ => [ (Blast_tac 1) ]);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   319
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   320
Addsimps [Sigma_empty1,Sigma_empty2]; 
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   321
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   322
goal Prod.thy "((a,b): Sigma A B) = (a:A & b:B(a))";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   323
by (Blast_tac 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   324
qed "mem_Sigma_iff";
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   325
AddIffs [mem_Sigma_iff]; 
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   326
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   327
val Collect_Prod = prove_goal Prod.thy 
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   328
	"{(a,b). P a & Q b} = Collect P Times Collect Q" (K [Blast_tac 1]);
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   329
Addsimps [Collect_Prod];
1515
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   330
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   331
(*Suggested by Pierre Chartier*)
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   332
goal Prod.thy
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   333
     "(UN (a,b):(A Times B). E a Times F b) = (UNION A E) Times (UNION B F)";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   334
by (Blast_tac 1);
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   335
qed "UNION_Times_distrib";
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   336
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   337
(*** Domain of a relation ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   338
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   339
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
   340
by (rtac CollectI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   341
by (rtac bexI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   342
by (rtac (fst_conv RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   343
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   344
qed "fst_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   345
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   346
val major::prems = goal Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   347
    "[| a : fst``r;  !!y.[| (a,y) : r |] ==> P |] ==> P"; 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   348
by (rtac (major RS imageE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   349
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   350
by (etac ssubst 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   351
by (rtac (surjective_pairing RS subst) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   352
by (assume_tac 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   353
qed "fst_imageE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   354
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   355
(*** Range of a relation ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   356
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   357
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
   358
by (rtac CollectI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   359
by (rtac bexI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   360
by (rtac (snd_conv RS sym) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   361
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   362
qed "snd_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   363
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   364
val major::prems = goal Prod.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   365
    "[| a : snd``r;  !!y.[| (y,a) : r |] ==> P |] ==> P"; 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   366
by (rtac (major RS imageE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   367
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   368
by (etac ssubst 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   369
by (rtac (surjective_pairing RS subst) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   370
by (assume_tac 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   371
qed "snd_imageE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   372
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   373
(** Exhaustion rule for unit -- a degenerate form of induction **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   374
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   375
goalw Prod.thy [Unity_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   376
    "u = ()";
2886
fd5645efa43d Now: unit = {True}
nipkow
parents: 2880
diff changeset
   377
by (stac (rewrite_rule [unit_def] Rep_unit RS singletonD RS sym) 1);
2880
a0fde30aa126 Removed (Unit) in Prod.
nipkow
parents: 2856
diff changeset
   378
by (rtac (Rep_unit_inverse RS sym) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   379
qed "unit_eq";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1746
diff changeset
   380
 
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1746
diff changeset
   381
AddIs  [fst_imageI, snd_imageI, prod_fun_imageI];
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   382
AddSEs [fst_imageE, snd_imageE, prod_fun_imageE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   383
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   384
structure Prod_Syntax =
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   385
struct
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   386
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   387
val unitT = Type("unit",[]);
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   388
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   389
fun mk_prod (T1,T2) = Type("*", [T1,T2]);
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   390
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   391
(*Maps the type T1*...*Tn to [T1,...,Tn], however nested*)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   392
fun factors (Type("*", [T1,T2])) = factors T1 @ factors T2
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   393
  | factors T                    = [T];
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   394
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   395
(*Make a correctly typed ordered pair*)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   396
fun mk_Pair (t1,t2) = 
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   397
  let val T1 = fastype_of t1
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   398
      and T2 = fastype_of t2
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   399
  in  Const("Pair", [T1, T2] ---> mk_prod(T1,T2)) $ t1 $ t2  end;
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   400
   
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   401
fun split_const(Ta,Tb,Tc) = 
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   402
    Const("split", [[Ta,Tb]--->Tc, mk_prod(Ta,Tb)] ---> Tc);
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   403
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   404
(*In ap_split S T u, term u expects separate arguments for the factors of S,
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   405
  with result type T.  The call creates a new term expecting one argument
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   406
  of type S.*)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   407
fun ap_split (Type("*", [T1,T2])) T3 u = 
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   408
      split_const(T1,T2,T3) $ 
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   409
      Abs("v", T1, 
2031
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   410
          ap_split T2 T3
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   411
             ((ap_split T1 (factors T2 ---> T3) (incr_boundvars 1 u)) $ 
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   412
              Bound 0))
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   413
  | ap_split T T3 u = u;
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   414
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   415
(*Makes a nested tuple from a list, following the product type structure*)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   416
fun mk_tuple (Type("*", [T1,T2])) tms = 
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   417
        mk_Pair (mk_tuple T1 tms, 
2031
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   418
                 mk_tuple T2 (drop (length (factors T1), tms)))
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   419
  | mk_tuple T (t::_) = t;
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   420
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   421
(*Attempts to remove occurrences of split, and pair-valued parameters*)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   422
val remove_split = rewrite_rule [split RS eq_reflection]  o  
2031
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   423
                   rule_by_tactic (ALLGOALS split_all_tac);
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   424
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   425
(*Uncurries any Var of function type in the rule*)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   426
fun split_rule_var (t as Var(v, Type("fun",[T1,T2])), rl) =
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   427
      let val T' = factors T1 ---> T2
2031
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   428
          val newt = ap_split T1 T2 (Var(v,T'))
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   429
          val cterm = Thm.cterm_of (#sign(rep_thm rl))
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   430
      in
2031
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   431
          remove_split (instantiate ([], [(cterm t, cterm newt)]) rl)
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   432
      end
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   433
  | split_rule_var (t,rl) = rl;
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   434
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   435
(*Uncurries ALL function variables occurring in a rule's conclusion*)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   436
fun split_rule rl = foldr split_rule_var (term_vars (concl_of rl), rl)
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   437
                    |> standard;
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   438
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   439
end;