src/HOL/Sum.ML
author paulson
Thu, 19 Aug 1999 15:11:12 +0200
changeset 7283 5cfe2944910a
parent 7254 fc7f95f293da
child 7293 959e060f4a2f
permissions -rw-r--r--
documented svc_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
     1
(*  Title:      HOL/Sum.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
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
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
     6
The disjoint sum of two types
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
(** Inl_Rep and Inr_Rep: Representations of the constructors **)
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*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    12
Goalw [Sum_def] "Inl_Rep(a) : Sum";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
by (EVERY1 [rtac CollectI, rtac disjI1, rtac exI, rtac refl]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
qed "Inl_RepI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    16
Goalw [Sum_def] "Inr_Rep(b) : Sum";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
qed "Inr_RepI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    20
Goal "inj_on Abs_Sum Sum";
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4535
diff changeset
    21
by (rtac inj_on_inverseI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
by (etac Abs_Sum_inverse 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4535
diff changeset
    23
qed "inj_on_Abs_Sum";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
(** Distinctness of Inl and Inr **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    27
Goalw [Inl_Rep_def, Inr_Rep_def] "Inl_Rep(a) ~= Inr_Rep(b)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
by (EVERY1 [rtac notI,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    29
            etac (fun_cong RS fun_cong RS fun_cong RS iffE), 
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    30
            rtac (notE RS ccontr),  etac (mp RS conjunct2), 
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    31
            REPEAT o (ares_tac [refl,conjI]) ]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
qed "Inl_Rep_not_Inr_Rep";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    34
Goalw [Inl_def,Inr_def] "Inl(a) ~= Inr(b)";
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4535
diff changeset
    35
by (rtac (inj_on_Abs_Sum RS inj_on_contraD) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
by (rtac Inl_Rep_not_Inr_Rep 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
by (rtac Inl_RepI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
by (rtac Inr_RepI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
qed "Inl_not_Inr";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
    41
bind_thm ("Inr_not_Inl", Inl_not_Inr RS not_sym);
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
    42
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
    43
AddIffs [Inl_not_Inr, Inr_not_Inl];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
    45
bind_thm ("Inl_neq_Inr", Inl_not_Inr RS notE);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
    47
val Inr_neq_Inl = sym RS Inl_neq_Inr;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
(** Injectiveness of Inl and Inr **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
    52
Goalw [Inl_Rep_def] "Inl_Rep(a) = Inl_Rep(c) ==> a=c";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
    53
by (etac (fun_cong RS fun_cong RS fun_cong RS iffE) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
    54
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
qed "Inl_Rep_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
    57
Goalw [Inr_Rep_def] "Inr_Rep(b) = Inr_Rep(d) ==> b=d";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
    58
by (etac (fun_cong RS fun_cong RS fun_cong RS iffE) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
    59
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
qed "Inr_Rep_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    62
Goalw [Inl_def] "inj(Inl)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
by (rtac injI 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4535
diff changeset
    64
by (etac (inj_on_Abs_Sum RS inj_onD RS Inl_Rep_inject) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
by (rtac Inl_RepI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
by (rtac Inl_RepI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
qed "inj_Inl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
val Inl_inject = inj_Inl RS injD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    70
Goalw [Inr_def] "inj(Inr)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
by (rtac injI 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4535
diff changeset
    72
by (etac (inj_on_Abs_Sum RS inj_onD RS Inr_Rep_inject) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
by (rtac Inr_RepI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
by (rtac Inr_RepI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
qed "inj_Inr";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
val Inr_inject = inj_Inr RS injD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    78
Goal "(Inl(x)=Inl(y)) = (x=y)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3842
diff changeset
    79
by (blast_tac (claset() addSDs [Inl_inject]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
qed "Inl_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
    82
Goal "(Inr(x)=Inr(y)) = (x=y)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3842
diff changeset
    83
by (blast_tac (claset() addSDs [Inr_inject]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
qed "Inr_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
    86
AddIffs [Inl_eq, Inr_eq];
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
    87
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
(*** Rules for the disjoint sum of two SETS ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
(** Introduction rules for the injections **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    92
Goalw [sum_def] "a : A ==> Inl(a) : A Plus B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
    93
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
qed "InlI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    96
Goalw [sum_def] "b : B ==> Inr(b) : A Plus B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
    97
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
qed "InrI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
(** Elimination rules **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
   102
val major::prems = Goalw [sum_def]
2212
bd705e9de196 plus -> Plus to avoid hiding class plus
nipkow
parents: 1985
diff changeset
   103
    "[| u: A Plus B;  \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
\       !!x. [| x:A;  u=Inl(x) |] ==> P; \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
\       !!y. [| y:B;  u=Inr(y) |] ==> P \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
by (rtac (major RS UnE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
by (REPEAT (rtac refl 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
     ORELSE eresolve_tac (prems@[imageE,ssubst]) 1));
2212
bd705e9de196 plus -> Plus to avoid hiding class plus
nipkow
parents: 1985
diff changeset
   110
qed "PlusE";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1515
diff changeset
   113
AddSIs [InlI, InrI]; 
2212
bd705e9de196 plus -> Plus to avoid hiding class plus
nipkow
parents: 1985
diff changeset
   114
AddSEs [PlusE];
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1515
diff changeset
   115
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
(** sum_case -- the selection operator for sums **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
7254
fc7f95f293da Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents: 7087
diff changeset
   119
Goalw [sum_case_def] "basic_sum_case f g (Inl x) = f(x)";
4535
f24cebc299e4 added select_equality to the implicit claset
oheimb
parents: 4477
diff changeset
   120
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
qed "sum_case_Inl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
7254
fc7f95f293da Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents: 7087
diff changeset
   123
Goalw [sum_case_def] "basic_sum_case f g (Inr x) = g(x)";
4535
f24cebc299e4 added select_equality to the implicit claset
oheimb
parents: 4477
diff changeset
   124
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
qed "sum_case_Inr";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
   127
Addsimps [sum_case_Inl, sum_case_Inr];
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
   128
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
(** Exhaustion rule for sums -- a degenerate form of induction **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
   131
val prems = Goalw [Inl_def,Inr_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
    "[| !!x::'a. s = Inl(x) ==> P;  !!y::'b. s = Inr(y) ==> P \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
by (rtac (rewrite_rule [Sum_def] Rep_Sum RS CollectE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
by (REPEAT (eresolve_tac [disjE,exE] 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
     ORELSE EVERY1 [resolve_tac prems, 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
   137
                    etac subst,
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
   138
                    rtac (Rep_Sum_inverse RS sym)]));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
qed "sumE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
   141
val prems = Goal "[| !!x. P (Inl x); !!x. P (Inr x) |] ==> P x";
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5148
diff changeset
   142
by (res_inst_tac [("s","x")] sumE 1);
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5148
diff changeset
   143
by (ALLGOALS (hyp_subst_tac THEN' (resolve_tac prems)));
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5148
diff changeset
   144
qed "sum_induct";
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5148
diff changeset
   145
7254
fc7f95f293da Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents: 7087
diff changeset
   146
Goal "basic_sum_case (%x::'a. f(Inl x)) (%y::'b. f(Inr y)) s = f(s)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
by (EVERY1 [res_inst_tac [("s","s")] sumE, 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
   148
            etac ssubst, rtac sum_case_Inl,
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
   149
            etac ssubst, rtac sum_case_Inr]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
qed "surjective_sum";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
7254
fc7f95f293da Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents: 7087
diff changeset
   152
Goal "R(basic_sum_case f g s) = \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
\             ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))";
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1761
diff changeset
   154
by (res_inst_tac [("s","s")] sumE 1);
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4089
diff changeset
   155
by Auto_tac;
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4535
diff changeset
   156
qed "split_sum_case";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   157
7254
fc7f95f293da Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents: 7087
diff changeset
   158
Goal "P (basic_sum_case f g s) = \
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7014
diff changeset
   159
\     (~ ((? x. s = Inl x & ~P (f x)) | (? y. s = Inr y & ~P (g y))))";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7014
diff changeset
   160
by (stac split_sum_case 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7014
diff changeset
   161
by (Blast_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7014
diff changeset
   162
qed "split_sum_case_asm";
4988
8f4dc836a2ea added split_sum_case_asm
oheimb
parents: 4830
diff changeset
   163
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
(*Prevents simplification of f and g: much faster*)
7254
fc7f95f293da Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents: 7087
diff changeset
   165
Goal "s=t ==> basic_sum_case f g s = basic_sum_case f g t";
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7014
diff changeset
   166
by (etac arg_cong 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7014
diff changeset
   167
qed "sum_case_weak_cong";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
7087
paulson
parents: 7031
diff changeset
   169
val [p1,p2] = Goal
7254
fc7f95f293da Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents: 7087
diff changeset
   170
  "[| basic_sum_case f1 f2 = basic_sum_case g1 g2;  \
7087
paulson
parents: 7031
diff changeset
   171
\     [| f1 = g1; f2 = g2 |] ==> P |] \
paulson
parents: 7031
diff changeset
   172
\  ==> P";
paulson
parents: 7031
diff changeset
   173
by (rtac p2 1);
paulson
parents: 7031
diff changeset
   174
by (rtac ext 1);
paulson
parents: 7031
diff changeset
   175
by (cut_inst_tac [("x","Inl x")] (p1 RS fun_cong) 1);
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5316
diff changeset
   176
by (Asm_full_simp_tac 1);
7087
paulson
parents: 7031
diff changeset
   177
by (rtac ext 1);
paulson
parents: 7031
diff changeset
   178
by (cut_inst_tac [("x","Inr x")] (p1 RS fun_cong) 1);
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5316
diff changeset
   179
by (Asm_full_simp_tac 1);
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5316
diff changeset
   180
qed "sum_case_inject";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
(** Rules for the Part primitive **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   185
Goalw [Part_def] "[| a : A;  a=h(b) |] ==> a : Part A h";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
   186
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
qed "Part_eqI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
val PartI = refl RSN (2,Part_eqI);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
   191
val major::prems = Goalw [Part_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
    "[| a : Part A h;  !!z. [| a : A;  a=h(z) |] ==> P  \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
by (rtac (major RS IntE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
by (etac CollectE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
by (etac exE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
qed "PartE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
   200
AddIs  [Part_eqI];
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
   201
AddSEs [PartE];
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
   202
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
   203
Goalw [Part_def] "Part A h <= A";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
by (rtac Int_lower1 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
qed "Part_subset";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   207
Goal "A<=B ==> Part A h <= Part B h";
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
   208
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
qed "Part_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
1515
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1465
diff changeset
   211
val basic_monos = basic_monos @ [Part_mono];
4ed79ebab64d Introduced normalize_thm into HOL.ML
nipkow
parents: 1465
diff changeset
   212
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   213
Goalw [Part_def] "a : Part A h ==> a : A";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   214
by (etac IntD1 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
qed "PartD1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
   217
Goal "Part A (%x. x) = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2212
diff changeset
   218
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
qed "Part_id";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   220
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
   221
Goal "Part (A Int B) h = (Part A h) Int (Part B h)";
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
   222
by (Blast_tac 1);
1188
0443e4dc8511 Added Part_Int and Part_Collect for inductive definitions
lcp
parents: 923
diff changeset
   223
qed "Part_Int";
0443e4dc8511 Added Part_Int and Part_Collect for inductive definitions
lcp
parents: 923
diff changeset
   224
0443e4dc8511 Added Part_Int and Part_Collect for inductive definitions
lcp
parents: 923
diff changeset
   225
(*For inductive definitions*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4988
diff changeset
   226
Goal "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}";
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
   227
by (Blast_tac 1);
1188
0443e4dc8511 Added Part_Int and Part_Collect for inductive definitions
lcp
parents: 923
diff changeset
   228
qed "Part_Collect";