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