src/HOL/BCV/JVM.ML
author wenzelm
Fri, 01 Dec 2000 19:43:06 +0100
changeset 10569 e8346dad78e1
parent 10199 7b6f9d34f737
child 10797 028d22926a41
permissions -rw-r--r--
ignore quick_and_dirty for coind;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/BCV/JVM.thy
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     2
    ID:         $Id$
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     4
    Copyright   2000 TUM
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     5
*)
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     6
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     7
Addsimps [Let_def];
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     8
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     9
Addsimps [is_type_def];
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    10
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    11
Goalw [states_def,JVM.sl_def,Opt.sl_def,Err.sl_def,
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    12
       stk_esl_def,reg_sl_def,Product.esl_def,
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    13
       Listn.sl_def,upto_esl_def,JType.esl_def,Err.esl_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    14
 "states S maxs maxr == opt(err((Union {list n (types S) |n. n <= maxs}) <*>\
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    15
\                                list maxr (err(types S))))";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    16
by (Simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    17
qed "JVM_states_unfold";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    18
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    19
Goalw [JVM.le_def,JVM.sl_def,Opt.sl_def,Err.sl_def,
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    20
       stk_esl_def,reg_sl_def,Product.esl_def,
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    21
       Listn.sl_def,upto_esl_def,JType.esl_def,Err.esl_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    22
 "JVM.le S m n == \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    23
\ Opt.le(Err.le(Product.le(Listn.le(subtype S))(Listn.le(Err.le(subtype S)))))";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    24
by (Simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    25
qed "JVM_le_unfold";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    26
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    27
Goalw [wfis_def,wfi_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    28
 "[| wfis S md maxr bs; bs!p = Load n; p < size bs |] ==> n < maxr";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    29
by (auto_tac (claset(),simpset()addsimps [all_set_conv_all_nth]));
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    30
qed "wf_LoadD";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    31
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    32
Goalw [wfis_def,wfi_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    33
 "[| wfis S md maxr bs; bs!p = Store n; p < size bs |] ==> n < maxr";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    34
by (auto_tac (claset(),simpset()addsimps [all_set_conv_all_nth]));
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    35
qed "wf_StoreD";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    36
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    37
Goalw [wfis_def,wfi_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    38
 "[| wfis S md m bs; bs!p = New C; p < size bs |] ==> (C,Object):S^*";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    39
by (auto_tac (claset(),simpset()addsimps [all_set_conv_all_nth]));
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    40
qed "wf_NewD";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    41
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    42
Goalw [wfis_def,wfi_def,wf_mdict_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    43
 "[| wfis S md maxr bs; wf_mdict S md; bs!p = Invoke C m pT (Class R); \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    44
\    p < size bs |] ==> (R,Object):S^*";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    45
by (force_tac (claset(),simpset()addsimps [all_set_conv_all_nth]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    46
qed "wf_InvokeD";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    47
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    48
Goalw [wfis_def,wfi_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    49
 "[|wfis S md m bs; bs!p = Getfield (Class C) D; p < size bs|] ==> \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    50
\ (C,Object):S^*";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    51
by (auto_tac (claset(),simpset()addsimps [all_set_conv_all_nth]));
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    52
qed "wf_GetfieldD";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    53
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    54
Goal "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    55
by (Blast_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    56
qed "special_ex_swap_lemma";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    57
AddIffs [special_ex_swap_lemma];
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    58
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    59
Goalw [pres_type_def,list_def,step_def,JVM_states_unfold]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    60
 "[| wfis S md maxr bs; wf_mdict S md |] ==> \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    61
\ pres_type (step S maxs rT bs) (size bs) (states S maxs maxr)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    62
by (clarify_tac (claset() addSEs [option_map_in_optionI,lift_in_errI]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    63
by (asm_simp_tac (simpset()addsimps [exec_def] addsplits [instr.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    64
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    65
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    66
by (fast_tac (claset() addDs [refl RS nth_in] addSEs [wf_LoadD] addss
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    67
      (simpset() delsimps [is_type_def]addsplits [err.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    68
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    69
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    70
by (fast_tac (claset() addDs [set_update_subset_insert RS subsetD]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    71
               addSEs [wf_StoreD] addss (simpset() addsplits [list.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    72
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    73
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    74
by (fast_tac (claset() addss (simpset() addsplits [list.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    75
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    76
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    77
by (fast_tac (claset() addIs [wf_GetfieldD]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    78
              addss (simpset() addsplits [list.split,ty.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    79
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    80
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    81
by (fast_tac (claset() addss (simpset() addsplits [list.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    82
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    83
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    84
by (fast_tac (claset() addIs [wf_NewD] addss simpset()) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    85
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    86
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    87
by (DEPTH_SOLVE_1(rtac conjI 1 ORELSE
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    88
    Clarify_tac 1 THEN asm_full_simp_tac (simpset() addsimps [wf_InvokeD]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    89
      addsplits [list.split,ty.split,option.split]) 1));
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    90
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    91
by (fast_tac (claset() addss (simpset() addsplits [list.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    92
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    93
qed "exec_pres_type";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    94
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    95
Goalw [wti_is_stable_topless_def,wti_def,stable_def,step_def,
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    96
       option_map_def,lift_def,bounded_def,JVM_le_unfold]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    97
 "[| bounded (succs bs) (size bs) |] ==> \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    98
\ wti_is_stable_topless (JVM.le S maxs maxr) \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    99
\    (Some Err) \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   100
\    (step S maxs rT bs) \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   101
\    (wti S maxs rT bs) \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   102
\    (succs bs) (size bs) (states S maxs maxr)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   103
by (simp_tac (simpset() addsplits [option.split]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   104
by (simp_tac (simpset() addsplits [err.split]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   105
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   106
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   107
 by (Blast_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   108
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   109
by (EVERY1[etac allE, mp_tac]);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   110
by (rewrite_goals_tac [exec_def,succs_def]);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   111
by (split_tac [instr.split] 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   112
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   113
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   114
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   115
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split,err.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   116
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   117
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   118
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   119
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   120
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   121
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   122
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   123
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   124
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   125
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   126
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   127
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   128
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   129
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   130
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   131
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   132
by (Blast_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   133
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   134
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   135
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   136
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   137
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   138
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   139
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   140
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   141
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   142
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   143
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   144
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   145
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   146
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   147
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   148
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split,ty.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   149
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   150
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   151
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   152
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   153
by (Blast_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   154
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   155
by (Clarify_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   156
by (asm_full_simp_tac (simpset() addsplits [option.split,list.split]
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   157
                                addsimps [le_list_refl,le_err_refl] ) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   158
qed "wti_is_stable_topless";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   159
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   160
Goalw [mono_def,step_def,option_map_def,lift_def,
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   161
       JVM_states_unfold,JVM_le_unfold]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   162
 "[| wfis S md maxr bs; wf_mdict S md |] ==> \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   163
\ mono (JVM.le S maxs maxr) (step S maxs rT bs) (size bs) (states S maxs maxr)";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   164
by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   165
by (simp_tac (simpset() addsplits [option.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   166
by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   167
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   168
by (split_tac [err.split] 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   169
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   170
 by (Clarify_tac 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   171
by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   172
by (simp_tac (simpset() addsimps [exec_def] delsplits [split_if]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   173
by (split_tac [instr.split] 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   174
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   175
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   176
by (fast_tac (claset() addDs [le_listD] addSEs [wf_LoadD]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   177
       addss (simpset() addsplits [err.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   178
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   179
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   180
by (fast_tac (claset() addIs [list_update_le_cong,le_listD] addSEs [wf_StoreD]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   181
       addss (simpset() addsplits [list.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   182
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   183
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   184
by (Asm_simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   185
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   186
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   187
by (Asm_simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   188
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   189
by (rtac conjI 1);
10199
7b6f9d34f737 delrules [r_into_rtrancl] required because the new I-rule made a step slow.
paulson
parents: 10172
diff changeset
   190
by (force_tac (claset(), simpset() addsplits [list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   191
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   192
by (rtac conjI 1);
10199
7b6f9d34f737 delrules [r_into_rtrancl] required because the new I-rule made a step slow.
paulson
parents: 10172
diff changeset
   193
by (fast_tac (claset() delrules [r_into_rtrancl]
7b6f9d34f737 delrules [r_into_rtrancl] required because the new I-rule made a step slow.
paulson
parents: 10172
diff changeset
   194
		       addDs [rtrancl_trans]
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   195
       addss (simpset() addsplits [list.split])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   196
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   197
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   198
 by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   199
 by (asm_full_simp_tac (simpset() addsplits [list.split]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   200
 by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   201
  by (Clarify_tac 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   202
 by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   203
 by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   204
  by (Clarify_tac 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   205
 by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   206
 by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   207
  by (Clarify_tac 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   208
 by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   209
 by (Asm_full_simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   210
 by (blast_tac (claset() addIs [rtrancl_trans,subtype_transD]) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   211
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   212
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   213
by (asm_full_simp_tac (simpset() addsimps [reflD] addsplits [list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   214
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   215
by (rtac conjI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   216
(* 39 *)
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   217
 by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   218
 by (asm_full_simp_tac (simpset() addsplits [list.split]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   219
 by (blast_tac (claset() addIs [rtrancl_trans,subtype_transD]) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   220
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   221
by (rtac conjI 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   222
by (fast_tac (claset() addss (simpset() addsplits [list.split,ty.split_asm]
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   223
                                       addsimps [is_Class_def,is_ref_def])) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   224
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   225
by (asm_simp_tac (simpset() addsplits [list.split]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   226
by (blast_tac (claset() addIs [subtype_transD]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   227
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   228
qed "exec_mono";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   229
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   230
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   231
Goalw [JVM.sl_def,stk_esl_def,reg_sl_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   232
 "[| univalent S; acyclic S |] ==> semilat(JVM.sl S maxs maxr)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   233
by (REPEAT(ares_tac [semilat_opt,err_semilat_Product_esl,err_semilat_upto_esl,
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   234
          err_semilat_eslI,semilat_Listn_sl,err_semilat_JType_esl] 1));
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   235
qed "semilat_JVM_slI";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   236
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   237
Goal "JVM.sl S maxs maxr == \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   238
\     (states S maxs maxr, JVM.le S maxs maxr, JVM.sup S maxs maxr)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   239
by (simp_tac (simpset() addsimps [states_def,JVM.le_def,JVM.sup_def,
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   240
      surjective_pairing RS sym]) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   241
qed "sl_triple_conv";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   242
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   243
Goalw [kiljvm_def,sl_triple_conv]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   244
 "[| univalent S; wf(S^-1); wfis S md maxr bs; wf_mdict S md; \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   245
\    bounded (succs bs) (size bs) |] ==> \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   246
\ is_bcv (JVM.le S maxs maxr) (Some Err) (wti S maxs rT bs) \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   247
\        (size bs) (states S maxs maxr) (kiljvm S maxs maxr rT bs)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   248
by (rtac is_bcv_kildall 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   249
      by (simp_tac (simpset() addsimps [symmetric sl_triple_conv]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   250
      by (blast_tac (claset() addSIs [semilat_JVM_slI] addDs [wf_acyclic]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   251
     by (simp_tac (simpset() addsimps [JVM_le_unfold]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   252
     by (blast_tac (claset()
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   253
       addSIs [acyclic_impl_order_subtype,wf_converse_subcls1_impl_acc_subtype]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   254
       addDs [wf_acyclic]) 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   255
    by (simp_tac (simpset() addsimps [JVM_le_unfold]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   256
   by (etac exec_pres_type 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   257
   by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   258
  by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   259
 by (etac exec_mono 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   260
 by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   261
by (etac wti_is_stable_topless 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   262
qed "is_bcv_kiljvm";