src/HOLCF/Up1.ML
author paulson
Thu, 29 Jun 2000 12:15:08 +0200
changeset 9187 68ecc04785f1
parent 9169 85a47aa21f74
child 9245 428385c4bc50
permissions -rw-r--r--
fixed proof to cope with the default of equalityCE instead of equalityE
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
     1
(*  Title:      HOLCF/Up1.ML
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
     2
    ID:         $Id$
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
     3
    Author:     Franz Regensburger
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
     4
    Copyright   1993  Technische Universitaet Muenchen
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
     5
*)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
     6
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     7
Goal "Rep_Up (Abs_Up y) = y";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     8
by (simp_tac (simpset() addsimps [Up_def,Abs_Up_inverse]) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
     9
qed "Abs_Up_inverse2";
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    10
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    11
qed_goalw "Exh_Up" thy [Iup_def ]
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    12
        "z = Abs_Up(Inl ()) | (? x. z = Iup x)"
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    13
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    14
        [
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    15
        (rtac (Rep_Up_inverse RS subst) 1),
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    16
        (res_inst_tac [("s","Rep_Up z")] sumE 1),
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    17
        (rtac disjI1 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    18
        (res_inst_tac [("f","Abs_Up")] arg_cong 1),
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    19
        (rtac (unit_eq RS subst) 1),
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    20
        (atac 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    21
        (rtac disjI2 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    22
        (rtac exI 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    23
        (res_inst_tac [("f","Abs_Up")] arg_cong 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    24
        (atac 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    25
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    26
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    27
Goal "inj(Abs_Up)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    28
by (rtac inj_inverseI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    29
by (rtac Abs_Up_inverse2 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    30
qed "inj_Abs_Up";
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    31
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    32
Goal "inj(Rep_Up)";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    33
by (rtac inj_inverseI 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    34
by (rtac Rep_Up_inverse 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    35
qed "inj_Rep_Up";
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    36
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    37
qed_goalw "inject_Iup" thy [Iup_def] "Iup x=Iup y ==> x=y"
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    38
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    39
        [
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    40
        (cut_facts_tac prems 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    41
        (rtac (inj_Inr RS injD) 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    42
        (rtac (inj_Abs_Up RS injD) 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    43
        (atac 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    44
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    45
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    46
AddSDs [inject_Iup];
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    47
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    48
qed_goalw "defined_Iup" thy [Iup_def] "Iup x~=Abs_Up(Inl ())"
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    49
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    50
        [
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    51
        (rtac notI 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    52
        (rtac notE 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    53
        (rtac Inl_not_Inr 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    54
        (rtac sym 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    55
        (etac (inj_Abs_Up RS  injD) 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    56
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    57
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    58
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    59
val prems = Goal "[| p=Abs_Up(Inl ()) ==> Q; !!x. p=Iup(x)==>Q|] ==>Q";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    60
by (rtac (Exh_Up RS disjE) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    61
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    62
by (etac exE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    63
by (eresolve_tac prems 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    64
qed "upE";
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    65
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    66
qed_goalw "Ifup1"  thy [Ifup_def]
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    67
        "Ifup(f)(Abs_Up(Inl ()))=UU"
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    68
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    69
        [
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    70
        (stac Abs_Up_inverse2 1),
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    71
        (stac sum_case_Inl 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    72
        (rtac refl 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    73
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    74
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    75
qed_goalw "Ifup2"  thy [Ifup_def,Iup_def]
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    76
        "Ifup(f)(Iup(x))=f`x"
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    77
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    78
        [
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    79
        (stac Abs_Up_inverse2 1),
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    80
        (stac sum_case_Inr 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    81
        (rtac refl 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    82
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    83
8161
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 7294
diff changeset
    84
val Up0_ss = (simpset_of Cfun3.thy) delsimps [range_composition] 
bde1391fd0a5 added range_composition (also to simpset)
oheimb
parents: 7294
diff changeset
    85
	     addsimps [Ifup1,Ifup2];
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    86
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    87
Addsimps [Ifup1,Ifup2];
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
    88
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    89
qed_goalw "less_up1a"  thy [less_up_def]
3323
194ae2e0c193 eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents: 2640
diff changeset
    90
        "Abs_Up(Inl ())<< z"
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    91
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    92
        [
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    93
        (stac Abs_Up_inverse2 1),
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    94
        (stac sum_case_Inl 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    95
        (rtac TrueI 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    96
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
    97
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
    98
qed_goalw "less_up1b" thy [Iup_def,less_up_def]
3323
194ae2e0c193 eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents: 2640
diff changeset
    99
        "~(Iup x) << (Abs_Up(Inl ()))"
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   100
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   101
        [
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   102
        (rtac notI 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   103
        (rtac iffD1 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   104
        (atac 2),
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
   105
        (stac Abs_Up_inverse2 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
   106
        (stac Abs_Up_inverse2 1),
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   107
        (stac sum_case_Inr 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   108
        (stac sum_case_Inl 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   109
        (rtac refl 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   110
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   111
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
   112
qed_goalw "less_up1c"  thy [Iup_def,less_up_def]
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   113
        "(Iup x) << (Iup y)=(x<<y)"
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   114
 (fn prems =>
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   115
        [
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
   116
        (stac Abs_Up_inverse2 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2278
diff changeset
   117
        (stac Abs_Up_inverse2 1),
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   118
        (stac sum_case_Inr 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   119
        (stac sum_case_Inr 1),
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   120
        (rtac refl 1)
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   121
        ]);
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   122
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   123
AddIffs [less_up1a, less_up1b, less_up1c];
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   124
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   125
Goal "(p::'a u) << p";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   126
by (res_inst_tac [("p","p")] upE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   127
by Auto_tac;
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   128
qed "refl_less_up";
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   129
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   130
Goal "[|(p1::'a u) << p2;p2 << p1|] ==> p1=p2";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   131
by (res_inst_tac [("p","p1")] upE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   132
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   133
by (res_inst_tac [("p","p2")] upE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   134
by (etac sym 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   135
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   136
by (res_inst_tac [("P","(Iup x) << (Abs_Up(Inl ()))")] notE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   137
by (rtac less_up1b 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   138
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   139
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   140
by (res_inst_tac [("p","p2")] upE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   141
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   142
by (res_inst_tac [("P","(Iup x) << (Abs_Up(Inl ()))")] notE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   143
by (rtac less_up1b 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   144
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   145
by (blast_tac (claset() addIs [arg_cong, antisym_less, less_up1c RS iffD1]) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   146
qed "antisym_less_up";
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   147
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   148
Goal "[|(p1::'a u) << p2;p2 << p3|] ==> p1 << p3";
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   149
by (res_inst_tac [("p","p1")] upE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   150
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   151
by (rtac less_up1a 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   152
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   153
by (res_inst_tac [("p","p2")] upE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   154
by (hyp_subst_tac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   155
by (rtac notE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   156
by (rtac less_up1b 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   157
by (atac 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   158
by (res_inst_tac [("p","p3")] upE 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   159
by Auto_tac;
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   160
by (blast_tac (claset() addIs [trans_less]) 1);
85a47aa21f74 tidying and unbatchifying
paulson
parents: 8161
diff changeset
   161
qed "trans_less_up";
2278
d63ffafce255 *** empty log message ***
oheimb
parents:
diff changeset
   162