src/HOLCF/Ssum0.ML
author nipkow
Tue, 09 Jan 2001 15:32:27 +0100
changeset 10834 a7897aebbffc
parent 10230 5eb935d6cc69
child 12030 46d57d0290a2
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     1
(*  Title:      HOLCF/Ssum0.ML
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1277
diff changeset
     3
    Author:     Franz Regensburger
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     4
    Copyright   1993  Technische Universitaet Muenchen
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     5
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     6
Strict sum with typedef
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     7
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     8
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     9
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    10
(* A non-emptyness result for Sssum                                         *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    11
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    12
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    13
Goalw [Ssum_def] "Sinl_Rep(a):Ssum";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
    14
by (Blast_tac 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    15
qed "SsumIl";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    16
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    17
Goalw [Ssum_def] "Sinr_Rep(a):Ssum";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
    18
by (Blast_tac 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    19
qed "SsumIr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    20
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    21
Goal "inj_on Abs_Ssum Ssum";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    22
by (rtac inj_on_inverseI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    23
by (etac Abs_Ssum_inverse 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    24
qed "inj_on_Abs_Ssum";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    25
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    26
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    27
(* Strictness of Sinr_Rep, Sinl_Rep and Isinl, Isinr                        *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    28
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    29
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    30
Goalw [Sinr_Rep_def,Sinl_Rep_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    31
 "Sinl_Rep(UU) = Sinr_Rep(UU)";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    32
by (rtac ext 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    33
by (rtac ext 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    34
by (rtac ext 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    35
by (fast_tac HOL_cs 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    36
qed "strict_SinlSinr_Rep";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    37
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    38
Goalw [Isinl_def,Isinr_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    39
 "Isinl(UU) = Isinr(UU)";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    40
by (rtac (strict_SinlSinr_Rep RS arg_cong) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    41
qed "strict_IsinlIsinr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    42
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    43
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    44
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    45
(* distinctness of  Sinl_Rep, Sinr_Rep and Isinl, Isinr                     *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    46
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    47
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    48
Goalw [Sinl_Rep_def,Sinr_Rep_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    49
        "(Sinl_Rep(a) = Sinr_Rep(b)) ==> a=UU & b=UU";
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    50
by (blast_tac (claset() addSDs [fun_cong]) 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    51
qed "noteq_SinlSinr_Rep";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    52
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    53
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    54
Goalw [Isinl_def,Isinr_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    55
        "Isinl(a)=Isinr(b) ==> a=UU & b=UU";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    56
by (rtac noteq_SinlSinr_Rep 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    57
by (etac (inj_on_Abs_Ssum  RS inj_onD) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    58
by (rtac SsumIl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    59
by (rtac SsumIr 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    60
qed "noteq_IsinlIsinr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    61
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    62
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    63
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    64
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    65
(* injectivity of Sinl_Rep, Sinr_Rep and Isinl, Isinr                       *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    66
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    67
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    68
Goalw [Sinl_Rep_def] "(Sinl_Rep(a) = Sinl_Rep(UU)) ==> a=UU";
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    69
by (blast_tac (claset() addSDs [fun_cong]) 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    70
qed "inject_Sinl_Rep1";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    71
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    72
Goalw [Sinr_Rep_def] "(Sinr_Rep(b) = Sinr_Rep(UU)) ==> b=UU";
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    73
by (blast_tac (claset() addSDs [fun_cong]) 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    74
qed "inject_Sinr_Rep1";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    75
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    76
Goalw [Sinl_Rep_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    77
"[| a1~=UU ; a2~=UU ; Sinl_Rep(a1)=Sinl_Rep(a2) |] ==> a1=a2";
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    78
by (blast_tac (claset() addSDs [fun_cong]) 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    79
qed "inject_Sinl_Rep2";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    80
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    81
Goalw [Sinr_Rep_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    82
"[|b1~=UU ; b2~=UU ; Sinr_Rep(b1)=Sinr_Rep(b2) |] ==> b1=b2";
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    83
by (blast_tac (claset() addSDs [fun_cong]) 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    84
qed "inject_Sinr_Rep2";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    85
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    86
Goal "Sinl_Rep(a1)=Sinl_Rep(a2) ==> a1=a2";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    87
by (case_tac "a1=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    88
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    89
by (rtac (inject_Sinl_Rep1 RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    90
by (etac sym 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    91
by (case_tac "a2=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    92
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    93
by (etac inject_Sinl_Rep1 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    94
by (etac inject_Sinl_Rep2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    95
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    96
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    97
qed "inject_Sinl_Rep";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    98
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    99
Goal "Sinr_Rep(b1)=Sinr_Rep(b2) ==> b1=b2";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   100
by (case_tac "b1=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   101
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   102
by (rtac (inject_Sinr_Rep1 RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   103
by (etac sym 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   104
by (case_tac "b2=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   105
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   106
by (etac inject_Sinr_Rep1 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   107
by (etac inject_Sinr_Rep2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   108
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   109
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   110
qed "inject_Sinr_Rep";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   111
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   112
Goalw [Isinl_def] "Isinl(a1)=Isinl(a2)==> a1=a2";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   113
by (rtac inject_Sinl_Rep 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   114
by (etac (inj_on_Abs_Ssum  RS inj_onD) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   115
by (rtac SsumIl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   116
by (rtac SsumIl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   117
qed "inject_Isinl";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   118
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   119
Goalw [Isinr_def] "Isinr(b1)=Isinr(b2) ==> b1=b2";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   120
by (rtac inject_Sinr_Rep 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   121
by (etac (inj_on_Abs_Ssum  RS inj_onD) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   122
by (rtac SsumIr 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   123
by (rtac SsumIr 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   124
qed "inject_Isinr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   125
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   126
AddSDs [inject_Isinl, inject_Isinr];
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   127
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   128
Goal "a1~=a2 ==> Isinl(a1) ~= Isinl(a2)";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   129
by (Blast_tac 1);
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   130
qed "inject_Isinl_rev";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   131
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   132
Goal "b1~=b2 ==> Isinr(b1) ~= Isinr(b2)";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   133
by (Blast_tac 1);
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   134
qed "inject_Isinr_rev";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   135
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   136
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   137
(* Exhaustion of the strict sum ++                                          *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   138
(* choice of the bottom representation is arbitrary                         *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   139
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   140
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   141
Goalw [Isinl_def,Isinr_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   142
        "z=Isinl(UU) | (? a. z=Isinl(a) & a~=UU) | (? b. z=Isinr(b) & b~=UU)";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   143
by (rtac (rewrite_rule [Ssum_def] Rep_Ssum RS CollectE) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   144
by (etac disjE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   145
by (etac exE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   146
by (case_tac "z= Abs_Ssum(Sinl_Rep(UU))" 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   147
by (etac disjI1 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   148
by (rtac disjI2 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   149
by (rtac disjI1 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   150
by (rtac exI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   151
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   152
by (rtac (Rep_Ssum_inverse RS sym RS trans) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   153
by (etac arg_cong 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   154
by (res_inst_tac [("Q","Sinl_Rep(a)=Sinl_Rep(UU)")] contrapos_nn 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   155
by (etac arg_cong 2);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   156
by (etac contrapos_nn 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   157
by (rtac (Rep_Ssum_inverse RS sym RS trans) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   158
by (rtac trans 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   159
by (etac arg_cong 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   160
by (etac arg_cong 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   161
by (etac exE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   162
by (case_tac "z= Abs_Ssum(Sinl_Rep(UU))" 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   163
by (etac disjI1 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   164
by (rtac disjI2 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   165
by (rtac disjI2 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   166
by (rtac exI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   167
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   168
by (rtac (Rep_Ssum_inverse RS sym RS trans) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   169
by (etac arg_cong 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   170
by (res_inst_tac [("Q","Sinr_Rep(b)=Sinl_Rep(UU)")] contrapos_nn 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   171
by (hyp_subst_tac 2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   172
by (rtac (strict_SinlSinr_Rep RS sym) 2);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   173
by (etac contrapos_nn 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   174
by (rtac (Rep_Ssum_inverse RS sym RS trans) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   175
by (rtac trans 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   176
by (etac arg_cong 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   177
by (etac arg_cong 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   178
qed "Exh_Ssum";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   179
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   180
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   181
(* elimination rules for the strict sum ++                                  *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   182
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   183
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   184
val prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1277
diff changeset
   185
        "[|p=Isinl(UU) ==> Q ;\
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1277
diff changeset
   186
\       !!x.[|p=Isinl(x); x~=UU |] ==> Q;\
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   187
\       !!y.[|p=Isinr(y); y~=UU |] ==> Q|] ==> Q";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   188
by (rtac (Exh_Ssum RS disjE) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   189
by (etac disjE 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   190
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   191
by (etac exE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   192
by (etac conjE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   193
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   194
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   195
by (etac exE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   196
by (etac conjE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   197
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   198
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   199
qed "IssumE";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   200
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   201
val prems = Goal
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   202
"[| !!x. [| p = Isinl(x) |] ==> Q;   !!y. [| p = Isinr(y) |] ==> Q |] ==>Q";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   203
by (rtac IssumE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   204
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   205
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   206
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   207
qed "IssumE2";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   208
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   209
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   210
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   211
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   212
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   213
(* rewrites for Iwhen                                                       *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   214
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   215
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   216
Goalw [Iwhen_def]
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   217
        "Iwhen f g (Isinl UU) = UU";
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9248
diff changeset
   218
by (rtac some_equality 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   219
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   220
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   221
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   222
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   223
by (res_inst_tac [("P","a=UU")] notE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   224
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   225
by (rtac inject_Isinl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   226
by (rtac sym 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   227
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   228
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   229
by (res_inst_tac [("P","b=UU")] notE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   230
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   231
by (rtac inject_Isinr 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   232
by (rtac sym 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   233
by (rtac (strict_IsinlIsinr RS subst) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   234
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   235
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   236
qed "Iwhen1";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   237
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   238
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   239
Goalw [Iwhen_def]
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   240
        "x~=UU ==> Iwhen f g (Isinl x) = f$x";
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9248
diff changeset
   241
by (rtac some_equality 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   242
by (fast_tac HOL_cs  2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   243
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   244
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   245
by (res_inst_tac [("P","x=UU")] notE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   246
by (atac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   247
by (rtac inject_Isinl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   248
by (atac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   249
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   250
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   251
by (rtac cfun_arg_cong 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   252
by (rtac inject_Isinl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   253
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   254
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   255
by (res_inst_tac [("P","Isinl(x) = Isinr(b)")] notE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   256
by (fast_tac HOL_cs  2);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   257
by (rtac contrapos_nn 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   258
by (etac noteq_IsinlIsinr 2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   259
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   260
qed "Iwhen2";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   261
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   262
Goalw [Iwhen_def]
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   263
        "y~=UU ==> Iwhen f g (Isinr y) = g$y";
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9248
diff changeset
   264
by (rtac some_equality 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   265
by (fast_tac HOL_cs  2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   266
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   267
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   268
by (res_inst_tac [("P","y=UU")] notE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   269
by (atac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   270
by (rtac inject_Isinr 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   271
by (rtac (strict_IsinlIsinr RS subst) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   272
by (atac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   273
by (rtac conjI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   274
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   275
by (res_inst_tac [("P","Isinr(y) = Isinl(a)")] notE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   276
by (fast_tac HOL_cs  2);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 9969
diff changeset
   277
by (rtac contrapos_nn 1);
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   278
by (etac (sym RS noteq_IsinlIsinr) 2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   279
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   280
by (strip_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   281
by (rtac cfun_arg_cong 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   282
by (rtac inject_Isinr 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   283
by (fast_tac HOL_cs  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   284
qed "Iwhen3";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   285
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   286
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   287
(* instantiate the simplifier                                               *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   288
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   289
8161
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 4833
diff changeset
   290
val Ssum0_ss = (simpset_of Cfun3.thy) delsimps [range_composition] addsimps 
1277
caef3601c0b2 corrected some errors that occurred after introduction of local simpsets
regensbu
parents: 1274
diff changeset
   291
                [(strict_IsinlIsinr RS sym),Iwhen1,Iwhen2,Iwhen3];
caef3601c0b2 corrected some errors that occurred after introduction of local simpsets
regensbu
parents: 1274
diff changeset
   292
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   293
Addsimps [strict_IsinlIsinr RS sym, Iwhen1, Iwhen2, Iwhen3];