src/HOLCF/Ssum3.ML
author kleing
Mon, 21 Jun 2004 10:25:57 +0200
changeset 14981 e73f8140af78
parent 12030 46d57d0290a2
child 15568 41bfe19eabe2
permissions -rw-r--r--
Merged in license change from Isabelle2004
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     1
(*  Title:      HOLCF/Ssum3.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
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     5
Class instance of  ++ for class pcpo
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     6
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     7
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
     8
(* for compatibility with old HOLCF-Version *)
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     9
Goal "UU = Isinl UU";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    10
by (simp_tac (HOL_ss addsimps [UU_def,UU_ssum_def]) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    11
qed "inst_ssum_pcpo";
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    12
10198
paulson
parents: 9248
diff changeset
    13
Addsimps [inst_ssum_pcpo RS sym];
paulson
parents: 9248
diff changeset
    14
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    15
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    16
(* continuity for Isinl and Isinr                                           *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    17
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    18
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    19
Goal "contlub(Isinl)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    20
by (rtac contlubI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    21
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    22
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    23
by (rtac (thelub_ssum1a RS sym) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    24
by (rtac allI 3);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    25
by (rtac exI 3);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    26
by (rtac refl 3);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    27
by (etac (monofun_Isinl RS ch2ch_monofun) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    28
by (case_tac "lub(range(Y))=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    29
by (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    30
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    31
by (res_inst_tac [("f","Isinl")] arg_cong  1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    32
by (rtac (chain_UU_I_inverse RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    33
by (rtac allI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    34
by (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    35
by (etac (chain_UU_I RS spec ) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    36
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    37
by (rtac Iwhen1 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    38
by (res_inst_tac [("f","Isinl")] arg_cong  1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    39
by (rtac lub_equal 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    40
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    41
by (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    42
by (etac (monofun_Isinl RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    43
by (rtac allI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    44
by (case_tac "Y(k)=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    45
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    46
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    47
qed "contlub_Isinl";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    48
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    49
Goal "contlub(Isinr)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    50
by (rtac contlubI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    51
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    52
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    53
by (rtac (thelub_ssum1b RS sym) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    54
by (rtac allI 3);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    55
by (rtac exI 3);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    56
by (rtac refl 3);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    57
by (etac (monofun_Isinr RS ch2ch_monofun) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    58
by (case_tac "lub(range(Y))=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    59
by (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    60
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    61
by ((rtac arg_cong 1) THEN (rtac (chain_UU_I_inverse RS sym) 1));
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    62
by (rtac allI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    63
by (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    64
by (etac (chain_UU_I RS spec ) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    65
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    66
by (rtac (strict_IsinlIsinr RS subst) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    67
by (rtac Iwhen1 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    68
by ((rtac arg_cong 1) THEN (rtac lub_equal 1));
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    69
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    70
by (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    71
by (etac (monofun_Isinr RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    72
by (rtac allI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    73
by (case_tac "Y(k)=UU" 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    74
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    75
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    76
qed "contlub_Isinr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    77
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    78
Goal "cont(Isinl)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    79
by (rtac monocontlub2cont 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    80
by (rtac monofun_Isinl 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    81
by (rtac contlub_Isinl 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    82
qed "cont_Isinl";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    83
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    84
Goal "cont(Isinr)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    85
by (rtac monocontlub2cont 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    86
by (rtac monofun_Isinr 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    87
by (rtac contlub_Isinr 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    88
qed "cont_Isinr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    89
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    90
AddIffs [cont_Isinl, cont_Isinr];
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
    91
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    92
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    93
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    94
(* continuity for Iwhen in the firts two arguments                          *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    95
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    96
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    97
Goal "contlub(Iwhen)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    98
by (rtac contlubI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    99
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   100
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   101
by (rtac (thelub_fun RS sym) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   102
by (etac (monofun_Iwhen1 RS ch2ch_monofun) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   103
by (rtac (expand_fun_eq RS iffD2) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   104
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   105
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   106
by (rtac (thelub_fun RS sym) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   107
by (rtac ch2ch_fun 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   108
by (etac (monofun_Iwhen1 RS ch2ch_monofun) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   109
by (rtac (expand_fun_eq RS iffD2) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   110
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   111
by (res_inst_tac [("p","xa")] IssumE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   112
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   113
by (rtac (lub_const RS thelubI RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   114
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   115
by (etac contlub_cfun_fun 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   116
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   117
by (rtac (lub_const RS thelubI RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   118
qed "contlub_Iwhen1";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   119
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   120
Goal "contlub(Iwhen(f))";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   121
by (rtac contlubI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   122
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   123
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   124
by (rtac (thelub_fun RS sym) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   125
by (etac (monofun_Iwhen2 RS ch2ch_monofun) 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   126
by (rtac (expand_fun_eq RS iffD2) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   127
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   128
by (res_inst_tac [("p","x")] IssumE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   129
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   130
by (rtac (lub_const RS thelubI RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   131
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   132
by (rtac (lub_const RS thelubI RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   133
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   134
by (etac contlub_cfun_fun 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   135
qed "contlub_Iwhen2";
243
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
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   138
(* continuity for Iwhen in its third argument                               *)
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
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   141
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   142
(* first 5 ugly lemmas                                                      *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   143
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   144
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   145
Goal "[| chain(Y); lub(range(Y)) = Isinl(x)|] ==> !i.? x. Y(i)=Isinl(x)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   146
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   147
by (res_inst_tac [("p","Y(i)")] IssumE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   148
by (etac exI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   149
by (etac exI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   150
by (res_inst_tac [("P","y=UU")] notE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   151
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   152
by (rtac (less_ssum3d RS iffD1) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   153
by (etac subst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   154
by (etac subst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   155
by (etac is_ub_thelub 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   156
qed "ssum_lemma9";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   157
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   158
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   159
Goal "[| chain(Y); lub(range(Y)) = Isinr(x)|] ==> !i.? x. Y(i)=Isinr(x)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   160
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   161
by (res_inst_tac [("p","Y(i)")] IssumE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   162
by (rtac exI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   163
by (etac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   164
by (rtac strict_IsinlIsinr 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   165
by (etac exI 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   166
by (res_inst_tac [("P","xa=UU")] notE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   167
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   168
by (rtac (less_ssum3c RS iffD1) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   169
by (etac subst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   170
by (etac subst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   171
by (etac is_ub_thelub 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   172
qed "ssum_lemma10";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   173
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   174
Goal "[| chain(Y); lub(range(Y)) = Isinl(UU) |] ==>\
8161
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   175
\   Iwhen f g (lub(range Y)) = lub(range(%i. Iwhen f g (Y i)))";
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   176
by (asm_simp_tac Ssum0_ss 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   177
by (rtac (chain_UU_I_inverse RS sym) 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   178
by (rtac allI 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   179
by (res_inst_tac [("s","Isinl(UU)"),("t","Y(i)")] subst 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   180
by (rtac (inst_ssum_pcpo RS subst) 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   181
by (rtac (chain_UU_I RS spec RS sym) 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   182
by (atac 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   183
by (etac (inst_ssum_pcpo RS ssubst) 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   184
by (asm_simp_tac Ssum0_ss 1);
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 5439
diff changeset
   185
qed "ssum_lemma11";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   186
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   187
Goal "[| chain(Y); lub(range(Y)) = Isinl(x); x ~= UU |] ==>\
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   188
\   Iwhen f g (lub(range Y)) = lub(range(%i. Iwhen f g (Y i)))";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   189
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   190
by (res_inst_tac [("t","x")] subst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   191
by (rtac inject_Isinl 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   192
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   193
by (atac 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   194
by (rtac (thelub_ssum1a RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   195
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   196
by (etac ssum_lemma9 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   197
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   198
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   199
by (rtac contlub_cfun_arg 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   200
by (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   201
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   202
by (rtac lub_equal2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   203
by (rtac (chain_mono2 RS exE) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   204
by (atac 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   205
by (rtac chain_UU_I_inverse2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   206
by (stac inst_ssum_pcpo 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   207
by (etac contrapos_np 1);
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   208
by (rtac inject_Isinl 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   209
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   210
by (etac sym 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   211
by (etac notnotD 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   212
by (rtac exI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   213
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   214
by (rtac (ssum_lemma9 RS spec RS exE) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   215
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   216
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   217
by (res_inst_tac [("t","Y(i)")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   218
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   219
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   220
by (rtac cfun_arg_cong 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   221
by (rtac Iwhen2 1);
10198
paulson
parents: 9248
diff changeset
   222
by (Force_tac 1); 
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   223
by (res_inst_tac [("t","Y(i)")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   224
by (atac 1);
10198
paulson
parents: 9248
diff changeset
   225
by Auto_tac;  
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   226
by (stac Iwhen2 1);
10198
paulson
parents: 9248
diff changeset
   227
by (Force_tac 1); 
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   228
by (simp_tac (simpset_of Cfun3.thy) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   229
by (rtac (monofun_Rep_CFun2 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   230
by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   231
by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   232
qed "ssum_lemma12";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   233
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   234
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   235
Goal "[| chain(Y); lub(range(Y)) = Isinr(x); x ~= UU |] ==>\
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   236
\   Iwhen f g (lub(range Y)) = lub(range(%i. Iwhen f g (Y i)))";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   237
by (asm_simp_tac Ssum0_ss 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   238
by (res_inst_tac [("t","x")] subst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   239
by (rtac inject_Isinr 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   240
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   241
by (atac 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   242
by (rtac (thelub_ssum1b RS sym) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   243
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   244
by (etac ssum_lemma10 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   245
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   246
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   247
by (rtac contlub_cfun_arg 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   248
by (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   249
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   250
by (rtac lub_equal2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   251
by (rtac (chain_mono2 RS exE) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   252
by (atac 2);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   253
by (rtac chain_UU_I_inverse2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   254
by (stac inst_ssum_pcpo 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   255
by (etac contrapos_np 1);
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   256
by (rtac inject_Isinr 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   257
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   258
by (etac sym 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   259
by (rtac (strict_IsinlIsinr RS subst) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   260
by (etac notnotD 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   261
by (rtac exI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   262
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   263
by (rtac (ssum_lemma10 RS spec RS exE) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   264
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   265
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   266
by (res_inst_tac [("t","Y(i)")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   267
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   268
by (rtac trans 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   269
by (rtac cfun_arg_cong 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   270
by (rtac Iwhen3 1);
10198
paulson
parents: 9248
diff changeset
   271
by (Force_tac 1); 
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   272
by (res_inst_tac [("t","Y(i)")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   273
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   274
by (stac Iwhen3 1);
10198
paulson
parents: 9248
diff changeset
   275
by (Force_tac 1); 
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   276
by (res_inst_tac [("t","Y(i)")] ssubst 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   277
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   278
by (simp_tac (simpset_of Cfun3.thy) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   279
by (rtac (monofun_Rep_CFun2 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   280
by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   281
by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   282
qed "ssum_lemma13";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   283
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   284
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   285
Goal "contlub(Iwhen(f)(g))";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   286
by (rtac contlubI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   287
by (strip_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   288
by (res_inst_tac [("p","lub(range(Y))")] IssumE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   289
by (etac ssum_lemma11 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   290
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   291
by (etac ssum_lemma12 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   292
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   293
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   294
by (etac ssum_lemma13 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   295
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   296
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   297
qed "contlub_Iwhen3";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   298
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   299
Goal "cont(Iwhen)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   300
by (rtac monocontlub2cont 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   301
by (rtac monofun_Iwhen1 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   302
by (rtac contlub_Iwhen1 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   303
qed "cont_Iwhen1";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   304
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   305
Goal "cont(Iwhen(f))";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   306
by (rtac monocontlub2cont 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   307
by (rtac monofun_Iwhen2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   308
by (rtac contlub_Iwhen2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   309
qed "cont_Iwhen2";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   310
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   311
Goal "cont(Iwhen(f)(g))";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   312
by (rtac monocontlub2cont 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   313
by (rtac monofun_Iwhen3 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   314
by (rtac contlub_Iwhen3 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   315
qed "cont_Iwhen3";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   316
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   317
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   318
(* continuous versions of lemmas for 'a ++ 'b                               *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   319
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   320
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   321
Goalw [sinl_def] "sinl$UU =UU";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   322
by (simp_tac (Ssum0_ss addsimps [cont_Isinl]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   323
by (rtac (inst_ssum_pcpo RS sym) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   324
qed "strict_sinl";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   325
Addsimps [strict_sinl];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   326
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   327
Goalw [sinr_def] "sinr$UU=UU";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   328
by (simp_tac (Ssum0_ss addsimps [cont_Isinr]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   329
by (rtac (inst_ssum_pcpo RS sym) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   330
qed "strict_sinr";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   331
Addsimps [strict_sinr];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   332
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   333
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   334
        "sinl$a=sinr$b ==> a=UU & b=UU";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   335
by (auto_tac (claset() addSDs [noteq_IsinlIsinr], simpset()));
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   336
qed "noteq_sinlsinr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   337
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   338
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   339
        "sinl$a1=sinl$a2==> a1=a2";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   340
by Auto_tac;
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   341
qed "inject_sinl";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   342
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   343
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   344
        "sinr$a1=sinr$a2==> a1=a2";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   345
by Auto_tac;
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   346
qed "inject_sinr";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   347
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   348
AddSDs [inject_sinl, inject_sinr];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   349
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   350
Goal "x~=UU ==> sinl$x ~= UU";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   351
by (etac contrapos_nn 1);
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   352
by (rtac inject_sinl 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   353
by Auto_tac;
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   354
qed "defined_sinl";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   355
Addsimps [defined_sinl];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   356
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   357
Goal "x~=UU ==> sinr$x ~= UU";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   358
by (etac contrapos_nn 1);
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   359
by (rtac inject_sinr 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   360
by Auto_tac;
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   361
qed "defined_sinr";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   362
Addsimps [defined_sinr];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   363
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   364
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   365
        "z=UU | (? a. z=sinl$a & a~=UU) | (? b. z=sinr$b & b~=UU)";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   366
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   367
by (stac inst_ssum_pcpo 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   368
by (rtac Exh_Ssum 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   369
qed "Exh_Ssum1";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   370
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   371
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   372
val [major,prem2,prem3] = Goalw [sinl_def,sinr_def] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1277
diff changeset
   373
        "[|p=UU ==> Q ;\
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   374
\       !!x.[|p=sinl$x; x~=UU |] ==> Q;\
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   375
\       !!y.[|p=sinr$y; y~=UU |] ==> Q|] ==> Q";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   376
by (rtac (major RS IssumE) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   377
by (stac inst_ssum_pcpo 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   378
by (atac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   379
by (rtac prem2 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   380
by (atac 2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   381
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   382
by (rtac prem3 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   383
by (atac 2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   384
by (Asm_simp_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   385
qed "ssumE";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   386
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   387
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   388
val [preml,premr] = Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   389
      "[|!!x.[|p=sinl$x|] ==> Q;\
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   390
\        !!y.[|p=sinr$y|] ==> Q|] ==> Q";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   391
by (rtac IssumE2 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   392
by (rtac preml 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   393
by (rtac premr 2);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   394
by Auto_tac;  
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   395
qed "ssumE2";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   396
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   397
val tac = (REPEAT (resolve_tac (cont_lemmas1 @ [cont_Iwhen1,cont_Iwhen2,
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   398
                cont_Iwhen3,cont2cont_CF1L]) 1)); 
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   399
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   400
Goalw [sscase_def,sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   401
        "sscase$f$g$UU = UU";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   402
by (stac inst_ssum_pcpo 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   403
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   404
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   405
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   406
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   407
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   408
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   409
by (simp_tac Ssum0_ss  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   410
qed "sscase1";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   411
Addsimps [sscase1];
2566
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2033
diff changeset
   412
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2033
diff changeset
   413
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2033
diff changeset
   414
val tac = (REPEAT (resolve_tac (cont_lemmas1 @ [cont_Iwhen1,cont_Iwhen2,
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2033
diff changeset
   415
                cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1));
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   416
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   417
Goalw [sscase_def,sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   418
        "x~=UU==> sscase$f$g$(sinl$x) = f$x";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   419
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   420
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   421
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   422
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   423
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   424
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   425
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   426
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   427
by (asm_simp_tac Ssum0_ss  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   428
qed "sscase2";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   429
Addsimps [sscase2];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   430
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   431
Goalw [sscase_def,sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   432
        "x~=UU==> sscase$f$g$(sinr$x) = g$x";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   433
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   434
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   435
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   436
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   437
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   438
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   439
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   440
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   441
by (asm_simp_tac Ssum0_ss  1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   442
qed "sscase3";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   443
Addsimps [sscase3];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   444
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   445
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   446
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   447
        "(sinl$x << sinl$y) = (x << y)";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   448
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   449
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   450
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   451
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   452
by (rtac less_ssum3a 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   453
qed "less_ssum4a";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   454
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   455
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   456
        "(sinr$x << sinr$y) = (x << y)";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   457
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   458
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   459
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   460
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   461
by (rtac less_ssum3b 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   462
qed "less_ssum4b";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   463
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   464
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   465
        "(sinl$x << sinr$y) = (x = UU)";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   466
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   467
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   468
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   469
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   470
by (rtac less_ssum3c 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   471
qed "less_ssum4c";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   472
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   473
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   474
        "(sinr$x << sinl$y) = (x = UU)";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   475
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   476
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   477
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   478
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   479
by (rtac less_ssum3d 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   480
qed "less_ssum4d";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   481
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   482
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   483
        "chain(Y) ==> (!i.? x.(Y i)=sinl$x)|(!i.? y.(Y i)=sinr$y)";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   484
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   485
by (etac ssum_lemma4 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   486
qed "ssum_chainE";
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   487
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   488
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   489
Goalw [sinl_def,sinr_def,sscase_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   490
"[| chain(Y); !i.? x. Y(i) = sinl$x |] ==>\ 
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   491
\   lub(range(Y)) = sinl$(lub(range(%i. sscase$(LAM x. x)$(LAM y. UU)$(Y i))))";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   492
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   493
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   494
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   495
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   496
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   497
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   498
by (stac (beta_cfun RS ext) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   499
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   500
by (rtac thelub_ssum1a 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   501
by (atac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   502
by (rtac allI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   503
by (etac allE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   504
by (etac exE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   505
by (rtac exI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   506
by (etac box_equals 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   507
by (rtac refl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   508
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinl]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   509
qed "thelub_ssum2a";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   510
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   511
Goalw [sinl_def,sinr_def,sscase_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   512
"[| chain(Y); !i.? x. Y(i) = sinr$x |] ==>\ 
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   513
\   lub(range(Y)) = sinr$(lub(range(%i. sscase$(LAM y. UU)$(LAM x. x)$(Y i))))";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   514
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   515
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   516
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   517
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   518
by (stac beta_cfun 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   519
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   520
by (stac (beta_cfun RS ext) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   521
by tac;
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   522
by (rtac thelub_ssum1b 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   523
by (atac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   524
by (rtac allI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   525
by (etac allE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   526
by (etac exE 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   527
by (rtac exI 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   528
by (etac box_equals 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   529
by (rtac refl 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   530
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   531
qed "thelub_ssum2b";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   532
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   533
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   534
        "[| chain(Y); lub(range(Y)) = sinl$x|] ==> !i.? x. Y(i)=sinl$x";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   535
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   536
by (etac ssum_lemma9 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   537
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   538
qed "thelub_ssum2a_rev";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   539
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
   540
Goalw [sinl_def,sinr_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   541
     "[| chain(Y); lub(range(Y)) = sinr$x|] ==> !i.? x. Y(i)=sinr$x";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   542
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   543
by (etac ssum_lemma10 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   544
by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 9169
diff changeset
   545
qed "thelub_ssum2b_rev";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   546
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   547
Goal "chain(Y) ==>\ 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   548
\   lub(range(Y)) = sinl$(lub(range(%i. sscase$(LAM x. x)$(LAM y. UU)$(Y i))))\
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   549
\ | lub(range(Y)) = sinr$(lub(range(%i. sscase$(LAM y. UU)$(LAM x. x)$(Y i))))";
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   550
by (rtac (ssum_chainE RS disjE) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   551
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   552
by (rtac disjI1 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   553
by (etac thelub_ssum2a 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   554
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   555
by (rtac disjI2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   556
by (etac thelub_ssum2b 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   557
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   558
qed "thelub_ssum3";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   559
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10230
diff changeset
   560
Goal "sscase$sinl$sinr$z=z";
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   561
by (res_inst_tac [("p","z")] ssumE 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   562
by Auto_tac;
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   563
qed "sscase4";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   564
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   565
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   566
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   567
(* install simplifier for Ssum                                              *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   568
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   569
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 1267
diff changeset
   570
val Ssum_rews = [strict_sinl,strict_sinr,defined_sinl,defined_sinr,
5439
2e0c18eedfd0 renamed sswhen to sscase
oheimb
parents: 5291
diff changeset
   571
                sscase1,sscase2,sscase3];