src/HOL/Prod.ML
author wenzelm
Tue, 03 Oct 2000 01:14:52 +0200
changeset 10131 546686f0a6fb
parent 10012 4961c73b5f60
permissions -rw-r--r--
range declared as syntax;
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
5810
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
     6
Ordered Pairs, the Cartesian product type, the unit type
923
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
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
     9
(** unit **)
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    10
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    11
Goalw [Unity_def]
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    12
    "u = ()";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    13
by (stac (rewrite_rule [unit_def] Rep_unit RS singletonD RS sym) 1);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    14
by (rtac (Rep_unit_inverse RS sym) 1);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    15
qed "unit_eq";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    16
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    17
(*simplification procedure for unit_eq.
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    18
  Cannot use this rule directly -- it loops!*)
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    19
local
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    20
  val unit_pat = Thm.cterm_of (Theory.sign_of (the_context ())) (Free ("x", HOLogic.unitT));
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    21
  val unit_meta_eq = standard (mk_meta_eq unit_eq);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    22
  fun proc _ _ t =
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    23
    if HOLogic.is_unit t then None
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    24
    else Some unit_meta_eq;
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    25
in
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    26
  val unit_eq_proc = Simplifier.mk_simproc "unit_eq" [unit_pat] proc;
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    27
end;
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    28
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    29
Addsimprocs [unit_eq_proc];
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    30
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    31
Goal "(!!x::unit. PROP P x) == PROP P ()";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    32
by (Simp_tac 1);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    33
qed "unit_all_eq1";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    34
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    35
Goal "(!!x::unit. PROP P) == PROP P";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    36
by (rtac triv_forall_equality 1);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    37
qed "unit_all_eq2";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    38
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    39
Goal "P () ==> P x";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    40
by (Simp_tac 1);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    41
qed "unit_induct";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    42
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    43
(*This rewrite counters the effect of unit_eq_proc on (%u::unit. f u),
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    44
  replacing it by f rather than by %u.f(). *)
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    45
Goal "(%u::unit. f()) = f";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    46
by (rtac ext 1);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    47
by (Simp_tac 1);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    48
qed "unit_abs_eta_conv";
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    49
Addsimps [unit_abs_eta_conv];
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    50
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    51
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    52
(** prod **)
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
    53
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
    54
Goalw [Prod_def] "Pair_Rep a b : Prod";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
by (EVERY1 [rtac CollectI, rtac exI, rtac exI, rtac refl]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
qed "ProdI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
    58
Goalw [Pair_Rep_def] "Pair_Rep a b = Pair_Rep a' b' ==> a=a' & b=b'";
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
    59
by (dtac (fun_cong RS fun_cong) 1);
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
    60
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
qed "Pair_Rep_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
    63
Goal "inj_on Abs_Prod Prod";
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4828
diff changeset
    64
by (rtac inj_on_inverseI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
by (etac Abs_Prod_inverse 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4828
diff changeset
    66
qed "inj_on_Abs_Prod";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
    68
val prems = Goalw [Pair_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    69
    "[| (a, b) = (a',b');  [| a=a';  b=b' |] ==> R |] ==> R";
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4828
diff changeset
    70
by (rtac (inj_on_Abs_Prod RS inj_onD RS Pair_Rep_inject RS conjE) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
by (REPEAT (ares_tac (prems@[ProdI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
qed "Pair_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
    74
Goal "((a,b) = (a',b')) = (a=a' & b=b')";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    75
by (blast_tac (claset() addSEs [Pair_inject]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
qed "Pair_eq";
3429
nipkow
parents: 2935
diff changeset
    77
AddIffs [Pair_eq];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    79
Goalw [fst_def] "fst (a,b) = a";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
    80
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
qed "fst_conv";
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    82
Goalw [snd_def] "snd (a,b) = b";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
    83
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
qed "snd_conv";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
    85
Addsimps [fst_conv, snd_conv];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    87
Goal "fst (x, y) = a ==> x = a";
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    88
by (Asm_full_simp_tac 1);
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    89
qed "fst_eqD";
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    90
Goal "snd (x, y) = a ==> y = a";
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    91
by (Asm_full_simp_tac 1);
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    92
qed "snd_eqD";
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
    93
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
    94
Goalw [Pair_def] "? x y. p = (x,y)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
by (EVERY1[etac exE, etac exE, rtac exI, rtac exI,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
    97
           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
    98
qed "PairE_lemma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
   100
val [prem] = Goal "[| !!x y. p = (x,y) ==> Q |] ==> Q";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
by (rtac (PairE_lemma RS exE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
by (REPEAT (eresolve_tac [prem,exE] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
qed "PairE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
4819
ef2e8e2a10e1 improved pair_tac to call prune_params_tac afterwards
oheimb
parents: 4799
diff changeset
   105
fun pair_tac s = EVERY' [res_inst_tac [("p",s)] PairE, hyp_subst_tac,
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   106
                         K prune_params_tac];
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   107
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   108
(* Do not add as rewrite rule: invalidates some proofs in IMP *)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   109
Goal "p = (fst(p),snd(p))";
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   110
by (pair_tac "p" 1);
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   111
by (Asm_simp_tac 1);
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   112
qed "surjective_pairing";
9345
2f5f6824f888 added (surjective_pairing RS sym) to simpset
oheimb
parents: 9020
diff changeset
   113
Addsimps [surjective_pairing RS sym];
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   114
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   115
Goal "? x y. z = (x, y)";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   116
by (rtac exI 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   117
by (rtac exI 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   118
by (rtac surjective_pairing 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   119
qed "surj_pair";
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   120
Addsimps [surj_pair];
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   121
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   122
5699
5b9a359e083c split_paired_all.ML;
wenzelm
parents: 5553
diff changeset
   123
bind_thm ("split_paired_all",
5b9a359e083c split_paired_all.ML;
wenzelm
parents: 5553
diff changeset
   124
  SplitPairedAll.rule (standard (surjective_pairing RS eq_reflection)));
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   125
bind_thms ("split_tupled_all", [split_paired_all, unit_all_eq2]);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   126
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   127
(*
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   128
Addsimps [split_paired_all] does not work with simplifier
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   129
because it also affects premises in congrence rules,
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   130
where is can lead to premises of the form !!a b. ... = ?P(a,b)
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   131
which cannot be solved by reflexivity.
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   132
*)
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   133
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   134
(* replace parameters of product type by individual component parameters *)
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   135
local
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   136
  fun exists_paired_all prem =   (* FIXME check deeper nesting of params!?! *)
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   137
    Library.exists (can HOLogic.dest_prodT o #2) (Logic.strip_params prem);
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   138
  val ss = HOL_basic_ss
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   139
    addsimps [split_paired_all, unit_all_eq2, unit_abs_eta_conv]
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   140
    addsimprocs [unit_eq_proc];
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   141
  val split_tac = full_simp_tac ss;
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   142
in
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   143
  val split_all_tac = SUBGOAL (fn (prem,i) =>
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   144
    if exists_paired_all prem then split_tac i else no_tac);
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   145
end;
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   146
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   147
claset_ref() := claset()
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   148
  addSWrapper ("split_all_tac", fn tac2 => split_all_tac ORELSE' tac2);
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   149
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   150
Goal "(!x. P x) = (!a b. P(a,b))";
4650
91af1ef45d68 added split_all_tac to claset()
oheimb
parents: 4534
diff changeset
   151
by (Fast_tac 1);
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   152
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
   153
Addsimps [split_paired_All];
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   154
(* 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
   155
5715
5fc697ad232b Added theorem prod_induct (needed for rep_datatype).
berghofe
parents: 5699
diff changeset
   156
bind_thm ("prod_induct",
5fc697ad232b Added theorem prod_induct (needed for rep_datatype).
berghofe
parents: 5699
diff changeset
   157
  allI RS (allI RS (split_paired_All RS iffD2)) RS spec);
5fc697ad232b Added theorem prod_induct (needed for rep_datatype).
berghofe
parents: 5699
diff changeset
   158
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   159
Goal "(? x. P x) = (? a b. P(a,b))";
4650
91af1ef45d68 added split_all_tac to claset()
oheimb
parents: 4534
diff changeset
   160
by (Fast_tac 1);
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   161
qed "split_paired_Ex";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   162
Addsimps [split_paired_Ex];
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   163
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   164
Goalw [split_def] "split c (a,b) = c a b";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   165
by (Simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
qed "split";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   167
Addsimps [split];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
7339
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   169
(*Subsumes the old split_Pair when f is the identity function*)
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   170
Goal "split (%x y. f(x,y)) = f";
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   171
by (rtac ext 1);
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   172
by (pair_tac "x" 1);
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   173
by (Simp_tac 1);
7339
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   174
qed "split_Pair_apply";
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   175
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   176
(*Can't be added to simpset: loops!*)
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   177
Goal "(SOME x. P x) = (SOME (a,b). P(a,b))";
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   178
by (simp_tac (simpset() addsimps [split_Pair_apply]) 1);
1b4d7a851b34 split_paired_Eps and lemmas
paulson
parents: 7031
diff changeset
   179
qed "split_paired_Eps";
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   180
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   181
Goal "!!s t. (s=t) = (fst(s)=fst(t) & snd(s)=snd(t))";
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   182
by (split_all_tac 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   183
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
qed "Pair_fst_snd_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
10012
4961c73b5f60 AddXIs [prod_eqI];
wenzelm
parents: 9969
diff changeset
   186
Goal "fst p = fst q ==> snd p = snd q ==> p = q";
4961c73b5f60 AddXIs [prod_eqI];
wenzelm
parents: 9969
diff changeset
   187
by (asm_simp_tac (simpset() addsimps [Pair_fst_snd_eq]) 1);
4961c73b5f60 AddXIs [prod_eqI];
wenzelm
parents: 9969
diff changeset
   188
qed "prod_eqI";
4961c73b5f60 AddXIs [prod_eqI];
wenzelm
parents: 9969
diff changeset
   189
AddXIs [prod_eqI];
4961c73b5f60 AddXIs [prod_eqI];
wenzelm
parents: 9969
diff changeset
   190
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
(*Prevents simplification of c: much faster*)
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
   192
Goal "p=q ==> split c p = split c q";
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
   193
by (etac arg_cong 1);
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   194
qed "split_weak_cong";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   196
Goal "(%(x,y). f(x,y)) = f";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   197
by (rtac ext 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   198
by (split_all_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   199
by (rtac split 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   200
qed "split_eta";
1655
5be64540f275 Added a number of lemmas
nipkow
parents: 1642
diff changeset
   201
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   202
val prems = Goal "(!!x y. f x y = g(x,y)) ==> (%(x,y). f x y) = g";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   203
by (asm_simp_tac (simpset() addsimps prems@[split_eta]) 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   204
qed "cond_split_eta";
5294
a84dd70e9925 replaced split_etas by split_eta_proc
oheimb
parents: 5278
diff changeset
   205
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   206
(*simplification procedure for cond_split_eta.
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   207
  using split_eta a rewrite rule is not general enough, and using
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   208
  cond_split_eta directly would render some existing proofs very inefficient.
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   209
  similarly for split_beta. *)
5294
a84dd70e9925 replaced split_etas by split_eta_proc
oheimb
parents: 5278
diff changeset
   210
local
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   211
  fun  Pair_pat k 0 (Bound m) = (m = k)
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   212
  |    Pair_pat k i (Const ("Pair",  _) $ Bound m $ t) = i > 0 andalso
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   213
                        m = k+i andalso Pair_pat k (i-1) t
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   214
  |    Pair_pat _ _ _ = false;
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   215
  fun no_args k i (Abs (_, _, t)) = no_args (k+1) i t
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   216
  |   no_args k i (t $ u) = no_args k i t andalso no_args k i u
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   217
  |   no_args k i (Bound m) = m < k orelse m > k+i
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   218
  |   no_args _ _ _ = true;
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   219
  fun split_pat tp i (Abs  (_,_,t)) = if tp 0 i t then Some (i,t) else None
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   220
  |   split_pat tp i (Const ("split", _) $ Abs (_, _, t)) = split_pat tp (i+1) t
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   221
  |   split_pat tp i _ = None;
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   222
  fun metaeq sg lhs rhs = mk_meta_eq (prove_goalw_cterm []
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   223
        (cterm_of sg (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs))))
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   224
        (K [simp_tac (HOL_basic_ss addsimps [cond_split_eta]) 1]));
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   225
  val sign = sign_of (the_context ());
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   226
  fun simproc name patstr = Simplifier.mk_simproc name
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   227
                [Thm.read_cterm sign (patstr, HOLogic.termT)];
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   228
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   229
  val beta_patstr = "split f z";
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   230
  val  eta_patstr = "split f";
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   231
  fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k+1) i t
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   232
  |   beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   233
                        (beta_term_pat k i t andalso beta_term_pat k i u)
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   234
  |   beta_term_pat k i t = no_args k i t;
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   235
  fun  eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   236
  |    eta_term_pat _ _ _ = false;
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   237
  fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t)
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   238
  |   subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   239
                              else (subst arg k i t $ subst arg k i u)
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   240
  |   subst arg k i t = t;
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   241
  fun beta_proc sg _ (s as Const ("split", _) $ Abs (_, _, t) $ arg) =
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   242
        (case split_pat beta_term_pat 1 t of
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   243
        Some (i,f) => Some (metaeq sg s (subst arg 0 i f))
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   244
        | None => None)
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   245
  |   beta_proc _ _ _ = None;
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   246
  fun eta_proc sg _ (s as Const ("split", _) $ Abs (_, _, t)) =
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   247
        (case split_pat eta_term_pat 1 t of
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   248
          Some (_,ft) => Some (metaeq sg s (let val (f $ arg) = ft in f end))
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   249
        | None => None)
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   250
  |   eta_proc _ _ _ = None;
5294
a84dd70e9925 replaced split_etas by split_eta_proc
oheimb
parents: 5278
diff changeset
   251
in
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   252
  val split_beta_proc = simproc "split_beta" beta_patstr beta_proc;
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   253
  val split_eta_proc  = simproc "split_eta"   eta_patstr  eta_proc;
5294
a84dd70e9925 replaced split_etas by split_eta_proc
oheimb
parents: 5278
diff changeset
   254
end;
a84dd70e9925 replaced split_etas by split_eta_proc
oheimb
parents: 5278
diff changeset
   255
7495
affcfd2830b7 added theorems fst_eqD and snd_eqD, added split_beta_proc, new split_eta_proc
oheimb
parents: 7339
diff changeset
   256
Addsimprocs [split_beta_proc,split_eta_proc];
5294
a84dd70e9925 replaced split_etas by split_eta_proc
oheimb
parents: 5278
diff changeset
   257
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   258
Goal "(%(x,y). P x y) z = P (fst z) (snd z)";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   259
by (stac surjective_pairing 1 THEN rtac split 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   260
qed "split_beta";
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   261
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   262
(*For use with split_tac and the simplifier*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   263
Goal "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
   264
by (stac surjective_pairing 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   265
by (stac split 1);
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   266
by (Blast_tac 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4828
diff changeset
   267
qed "split_split";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   268
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   269
(* could be done after split_tac has been speeded up significantly:
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4828
diff changeset
   270
simpset_ref() := simpset() addsplits [split_split];
3568
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   271
   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
   272
   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
   273
*)
36ff1ab12021 Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents: 3429
diff changeset
   274
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   275
Goal "R (split c p) = (~(? x y. p = (x,y) & (~R (c x y))))";
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4828
diff changeset
   276
by (stac split_split 1);
4435
41a7e4f0e957 added expand_split_asm
oheimb
parents: 4423
diff changeset
   277
by (Simp_tac 1);
41a7e4f0e957 added expand_split_asm
oheimb
parents: 4423
diff changeset
   278
qed "expand_split_asm";
41a7e4f0e957 added expand_split_asm
oheimb
parents: 4423
diff changeset
   279
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   280
(** split used as a logical connective or set former **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   281
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   282
(*These rules are for use with blast_tac.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   283
  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
   284
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   285
Goal "!!p. [| !!a b. p=(a,b) ==> c a b |] ==> split c p";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1515
diff changeset
   286
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
   287
by (Asm_simp_tac 1);
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   288
qed "splitI2";
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   289
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 7495
diff changeset
   290
Goal "!!p. [| !!a b. (a,b)=p ==> c a b x |] ==> split c p x";
f531589c9fc1 added various little lemmas
oheimb
parents: 7495
diff changeset
   291
by (split_all_tac 1);
f531589c9fc1 added various little lemmas
oheimb
parents: 7495
diff changeset
   292
by (Asm_simp_tac 1);
f531589c9fc1 added various little lemmas
oheimb
parents: 7495
diff changeset
   293
qed "splitI2'";
f531589c9fc1 added various little lemmas
oheimb
parents: 7495
diff changeset
   294
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   295
Goal "c a b ==> split c (a,b)";
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   296
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   297
qed "splitI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   298
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
   299
val prems = Goalw [split_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   300
    "[| 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
   301
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   302
qed "splitE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   303
8157
b1a458416c92 added splitE', also to claset
oheimb
parents: 8100
diff changeset
   304
val prems = Goalw [split_def]
b1a458416c92 added splitE', also to claset
oheimb
parents: 8100
diff changeset
   305
    "[| split c p z;  !!x y. [| p = (x,y);  c x y z |] ==> Q |] ==> Q";
b1a458416c92 added splitE', also to claset
oheimb
parents: 8100
diff changeset
   306
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
b1a458416c92 added splitE', also to claset
oheimb
parents: 8100
diff changeset
   307
qed "splitE'";
b1a458416c92 added splitE', also to claset
oheimb
parents: 8100
diff changeset
   308
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
   309
val major::prems = Goal
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   310
    "[| Q (split P z);  !!x y. [|z = (x, y); Q (P x y)|] ==> R  \
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   311
\    |] ==> R";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   312
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   313
by (rtac (split_beta RS subst) 1 THEN rtac major 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   314
qed "splitE2";
4134
5c6cb2a25816 added several theorems
oheimb
parents: 4089
diff changeset
   315
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   316
Goal "split R (a,b) ==> R a b";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   317
by (etac (split RS iffD1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   318
qed "splitD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   319
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   320
Goal "z: c a b ==> z: split c (a,b)";
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 972
diff changeset
   321
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   322
qed "mem_splitI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   323
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   324
Goal "!!p. [| !!a b. p=(a,b) ==> z: c a b |] ==> z: split c p";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1515
diff changeset
   325
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
   326
by (Asm_simp_tac 1);
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   327
qed "mem_splitI2";
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1301
diff changeset
   328
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
   329
val prems = Goalw [split_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   330
    "[| 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
   331
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   332
qed "mem_splitE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   333
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 7495
diff changeset
   334
AddSIs [splitI, splitI2, splitI2', mem_splitI, mem_splitI2];
8157
b1a458416c92 added splitE', also to claset
oheimb
parents: 8100
diff changeset
   335
AddSEs [splitE, splitE', mem_splitE];
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   336
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   337
Goal "(%u. ? x y. u = (x, y) & P (x, y)) = P";
9020
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   338
by (rtac ext 1);
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   339
by (Fast_tac 1);
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   340
qed "split_eta_SetCompr";
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   341
Addsimps [split_eta_SetCompr];
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   342
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   343
Goal "(%u. ? x y. u = (x, y) & P x y) = split P";
9020
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   344
br ext 1;
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   345
by (Fast_tac 1);
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   346
qed "split_eta_SetCompr2";
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   347
Addsimps [split_eta_SetCompr2];
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   348
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   349
(* allows simplifications of nested splits in case of independent predicates *)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   350
Goal "(%(a,b). P & Q a b) = (%ab. P & split Q ab)";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   351
by (rtac ext 1);
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   352
by (Blast_tac 1);
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   353
qed "split_part";
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   354
Addsimps [split_part];
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   355
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   356
Goal "(@(x',y'). x = x' & y = y') = (x,y)";
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   357
by (Blast_tac 1);
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   358
qed "Eps_split_eq";
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   359
Addsimps [Eps_split_eq];
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   360
(*
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   361
the following  would be slightly more general,
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   362
but cannot be used as rewrite rule:
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   363
### Cannot add premise as rewrite rule because it contains (type) unknowns:
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   364
### ?y = .x
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   365
Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)";
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
   366
by (rtac some_equality 1);
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   367
by ( Simp_tac 1);
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   368
by (split_all_tac 1);
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   369
by (Asm_full_simp_tac 1);
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   370
qed "Eps_split_eq";
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   371
*)
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   372
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   373
(*** prod_fun -- action of the product functor upon functions ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   374
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   375
Goalw [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
   376
by (rtac split 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   377
qed "prod_fun";
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4435
diff changeset
   378
Addsimps [prod_fun];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   379
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5144
diff changeset
   380
Goal "prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   381
by (rtac ext 1);
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   382
by (pair_tac "x" 1);
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4435
diff changeset
   383
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   384
qed "prod_fun_compose";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   385
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   386
Goal "prod_fun (%x. x) (%y. y) = (%z. z)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   387
by (rtac ext 1);
4828
ee3317896a47 improved split_all_tac significantly
oheimb
parents: 4819
diff changeset
   388
by (pair_tac "z" 1);
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4435
diff changeset
   389
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   390
qed "prod_fun_ident";
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4435
diff changeset
   391
Addsimps [prod_fun_ident];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   392
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
   393
Goal "(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
   394
by (rtac image_eqI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   395
by (rtac (prod_fun RS sym) 1);
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
   396
by (assume_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   397
qed "prod_fun_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   398
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
   399
val major::prems = Goal
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   400
    "[| 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
   401
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   402
by (rtac (major RS imageE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   403
by (res_inst_tac [("p","x")] PairE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   404
by (resolve_tac prems 1);
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   405
by (Blast_tac 2);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   406
by (blast_tac (claset() addIs [prod_fun]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   407
qed "prod_fun_imageE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   408
5788
e3a98a7c0634 Domain r, Range r replace fst``r, snd``r
paulson
parents: 5761
diff changeset
   409
AddIs  [prod_fun_imageI];
e3a98a7c0634 Domain r, Range r replace fst``r, snd``r
paulson
parents: 5761
diff changeset
   410
AddSEs [prod_fun_imageE];
e3a98a7c0634 Domain r, Range r replace fst``r, snd``r
paulson
parents: 5761
diff changeset
   411
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4435
diff changeset
   412
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   413
(*** Disjoint union of a family of sets - Sigma ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   414
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   415
Goalw [Sigma_def] "[| a:A;  b:B(a) |] ==> (a,b) : Sigma A B";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   416
by (REPEAT (ares_tac [singletonI,UN_I] 1));
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   417
qed "SigmaI";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   418
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   419
AddSIs [SigmaI];
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   420
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   421
(*The general elimination rule*)
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   422
val major::prems = Goalw [Sigma_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   423
    "[| c: Sigma A B;  \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   424
\       !!x y.[| x:A;  y:B(x);  c=(x,y) |] ==> P \
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   425
\    |] ==> P";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   426
by (cut_facts_tac [major] 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   427
by (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ;
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   428
qed "SigmaE";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   429
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   430
(** Elimination of (a,b):A*B -- introduces no eigenvariables **)
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   431
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   432
Goal "(a,b) : Sigma A B ==> a : A";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   433
by (etac SigmaE 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   434
by (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   435
qed "SigmaD1";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   436
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   437
Goal "(a,b) : Sigma A B ==> b : B(a)";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   438
by (etac SigmaE 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   439
by (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   440
qed "SigmaD2";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   441
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9359
diff changeset
   442
val [major,minor]= Goal
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   443
    "[| (a,b) : Sigma A B;    \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   444
\       [| a:A;  b:B(a) |] ==> P   \
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   445
\    |] ==> P";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   446
by (rtac minor 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   447
by (rtac (major RS SigmaD1) 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   448
by (rtac (major RS SigmaD2) 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   449
qed "SigmaE2";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   450
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   451
AddSEs [SigmaE2, SigmaE];
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   452
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5303
diff changeset
   453
val prems = Goal
1642
21db0cf9a1a4 Using new "Times" infix
paulson
parents: 1618
diff changeset
   454
    "[| 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
   455
by (cut_facts_tac prems 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   456
by (blast_tac (claset() addIs (prems RL [subsetD])) 1);
1515
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   457
qed "Sigma_mono";
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   458
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   459
Goal "Sigma {} B = {}";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   460
by (Blast_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   461
qed "Sigma_empty1";
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   462
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   463
Goal "A <*> {} = {}";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   464
by (Blast_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6830
diff changeset
   465
qed "Sigma_empty2";
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   466
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   467
Addsimps [Sigma_empty1,Sigma_empty2];
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   468
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   469
Goal "UNIV <*> UNIV = UNIV";
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   470
by Auto_tac;
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   471
qed "UNIV_Times_UNIV";
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   472
Addsimps [UNIV_Times_UNIV];
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   473
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   474
Goal "- (UNIV <*> A) = UNIV <*> (-A)";
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   475
by Auto_tac;
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   476
qed "Compl_Times_UNIV1";
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   477
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   478
Goal "- (A <*> UNIV) = (-A) <*> UNIV";
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   479
by Auto_tac;
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   480
qed "Compl_Times_UNIV2";
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   481
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   482
Addsimps [Compl_Times_UNIV1, Compl_Times_UNIV2];
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   483
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4989
diff changeset
   484
Goal "((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
   485
by (Blast_tac 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   486
qed "mem_Sigma_iff";
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   487
AddIffs [mem_Sigma_iff];
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   488
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   489
Goal "x:C ==> (A <*> C <= B <*> C) = (A <= B)";
6016
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   490
by (Blast_tac 1);
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   491
qed "Times_subset_cancel2";
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   492
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   493
Goal "x:C ==> (A <*> C = B <*> C) = (A = B)";
6016
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   494
by (blast_tac (claset() addEs [equalityE]) 1);
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   495
qed "Times_eq_cancel2";
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   496
9020
1056cbbaeb29 added split_eta_SetCompr2 (also to simpset), generalized SetCompr_Sigma_eq
oheimb
parents: 8703
diff changeset
   497
Goal "Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))";
8261
e4726ac0863a added split_eta_SetCompr, SetCompr_Sigma_eq
oheimb
parents: 8255
diff changeset
   498
by (Fast_tac 1);
e4726ac0863a added split_eta_SetCompr, SetCompr_Sigma_eq
oheimb
parents: 8255
diff changeset
   499
qed "SetCompr_Sigma_eq";
5810
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   500
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   501
(*** Complex rules for Sigma ***)
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   502
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   503
Goal "{(a,b). P a & Q b} = Collect P <*> Collect Q";
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   504
by (Blast_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   505
qed "Collect_split";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   506
4534
6932c3ae3912 added select_equality to the implicit claset
oheimb
parents: 4521
diff changeset
   507
Addsimps [Collect_split];
1515
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1485
diff changeset
   508
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   509
(*Suggested by Pierre Chartier*)
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   510
Goal "(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2886
diff changeset
   511
by (Blast_tac 1);
6830
f8aed3706af7 renamed UNION_... to UN_... (to fit the convention)
paulson
parents: 6394
diff changeset
   512
qed "UN_Times_distrib";
2856
cdb908486a96 Reorganization of how classical rules are installed
paulson
parents: 2637
diff changeset
   513
6016
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   514
Goal "(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))";
5810
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   515
by (Fast_tac 1);
6016
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   516
qed "split_paired_Ball_Sigma";
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   517
Addsimps [split_paired_Ball_Sigma];
5810
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   518
6016
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   519
Goal "(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))";
5810
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   520
by (Fast_tac 1);
6016
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   521
qed "split_paired_Bex_Sigma";
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   522
Addsimps [split_paired_Bex_Sigma];
5810
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   523
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   524
Goal "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   525
by (Blast_tac 1);
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   526
qed "Sigma_Un_distrib1";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   527
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   528
Goal "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   529
by (Blast_tac 1);
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   530
qed "Sigma_Un_distrib2";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   531
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   532
Goal "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   533
by (Blast_tac 1);
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   534
qed "Sigma_Int_distrib1";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   535
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   536
Goal "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   537
by (Blast_tac 1);
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   538
qed "Sigma_Int_distrib2";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   539
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   540
Goal "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   541
by (Blast_tac 1);
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   542
qed "Sigma_Diff_distrib1";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   543
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   544
Goal "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   545
by (Blast_tac 1);
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   546
qed "Sigma_Diff_distrib2";
829cae93f132 new TIMES/Sigma rules
paulson
parents: 5788
diff changeset
   547
6016
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   548
Goal "Sigma (Union X) B = (UN A:X. Sigma A B)";
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   549
by (Blast_tac 1);
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   550
qed "Sigma_Union";
797c76d6ff04 new (and generalized) theorems about Sigma/Times
paulson
parents: 5810
diff changeset
   551
8255
38f96394c099 new distributive laws
paulson
parents: 8157
diff changeset
   552
(*Non-dependent versions are needed to avoid the need for higher-order
38f96394c099 new distributive laws
paulson
parents: 8157
diff changeset
   553
  matching, especially when the rules are re-oriented*)
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   554
Goal "(A Un B) <*> C = (A <*> C) Un (B <*> C)";
8255
38f96394c099 new distributive laws
paulson
parents: 8157
diff changeset
   555
by (Blast_tac 1);
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   556
qed "Times_Un_distrib1";
8255
38f96394c099 new distributive laws
paulson
parents: 8157
diff changeset
   557
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   558
Goal "(A Int B) <*> C = (A <*> C) Int (B <*> C)";
8255
38f96394c099 new distributive laws
paulson
parents: 8157
diff changeset
   559
by (Blast_tac 1);
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   560
qed "Times_Int_distrib1";
8255
38f96394c099 new distributive laws
paulson
parents: 8157
diff changeset
   561
8703
816d8f6513be Times -> <*>
nipkow
parents: 8320
diff changeset
   562
Goal "(A - B) <*> C = (A <*> C) - (B <*> C)";
8255
38f96394c099 new distributive laws
paulson
parents: 8157
diff changeset
   563
by (Blast_tac 1);
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   564
qed "Times_Diff_distrib1";
5088
e4aa78d1312f New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents: 5083
diff changeset
   565
e4aa78d1312f New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents: 5083
diff changeset
   566
5096
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   567
(*Attempts to remove occurrences of split, and pair-valued parameters*)
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   568
val remove_split = rewrite_rule [split RS eq_reflection] o
5096
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   569
                   rule_by_tactic (TRYALL split_all_tac);
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   570
5096
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   571
local
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   572
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   573
(*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
   574
  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
   575
  of type S.*)
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   576
fun ap_split (Type ("*", [T1, T2])) T3 u =
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   577
      HOLogic.split_const (T1, T2, T3) $
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   578
      Abs("v", T1,
2031
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   579
          ap_split T2 T3
9359
a4b990838074 improved unit theory: unit_all_eq1, unit_all_eq2, split_all_tac;
wenzelm
parents: 9345
diff changeset
   580
             ((ap_split T1 (HOLogic.prodT_factors T2 ---> T3) (incr_boundvars 1 u)) $
2031
03a843f0f447 Ran expandshort
paulson
parents: 1754
diff changeset
   581
              Bound 0))
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   582
  | ap_split T T3 u = u;
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   583
5096
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   584
(*Curries any Var of function type in the rule*)
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   585
fun split_rule_var' (t as Var (v, Type ("fun", [T1, T2])), rl) =
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   586
      let val T' = HOLogic.prodT_factors T1 ---> T2
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   587
          val newt = ap_split T1 T2 (Var (v, T'))
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   588
          val cterm = Thm.cterm_of (#sign (rep_thm rl))
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   589
      in
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   590
          instantiate ([], [(cterm t, cterm newt)]) rl
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   591
      end
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   592
  | split_rule_var' (t, rl) = rl;
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   593
5096
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   594
in
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   595
5096
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   596
val split_rule_var = standard o remove_split o split_rule_var';
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   597
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   598
(*Curries ALL function variables occurring in a rule's conclusion*)
84b00be693b4 Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents: 5088
diff changeset
   599
fun split_rule rl = remove_split (foldr split_rule_var' (term_vars (concl_of rl), rl))
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   600
                    |> standard;
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   601
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1727
diff changeset
   602
end;