src/HOL/BCV/DFAimpl.ML
author paulson
Thu, 11 Nov 1999 10:25:17 +0100
changeset 8005 b64d86018785
parent 7961 422ac6888c7f
child 8423 3c19160b6432
permissions -rw-r--r--
new-style infix declaration for "image"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/BCV/DFAimpl.thy
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     4
    Copyright   1999 TUM
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     5
*)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     6
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     7
(** merges **)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     8
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     9
Goal "!sos. size(merges t ps sos) = size sos";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    10
by (induct_tac "ps" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    11
by (Auto_tac);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    12
qed_spec_mp "length_merges";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    13
Addsimps [length_merges];
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    14
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    15
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    16
 "!xs. xs : listsn n (option A) --> x : A --> (!p : set ps. p<n) --> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    17
\      semilat A --> merges x ps xs : listsn n (option A)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    18
by (induct_tac "ps" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    19
by (Auto_tac);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    20
qed_spec_mp "merges_preserves_type";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    21
Addsimps [merges_preserves_type];
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    22
(*AddSIs [merges_preserves_type];*)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    23
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    24
Goal "semilat A ==> !xs. xs : listsn n (option A) --> x:A --> (!p:set ps. p<n) \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    25
\ --> xs <= merges x ps xs";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    26
by (induct_tac "ps" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    27
 by (Simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    28
by (Simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    29
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    30
by (rtac order_trans 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    31
 by (etac list_update_incr 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    32
   by (Blast_tac 2);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    33
  by (assume_tac 2);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    34
 by (etac (Some_in_option RS iffD2) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    35
by (blast_tac (claset() addSIs [listsnE_set]
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    36
          addIs [semilat_plus,listsnE_length RS nth_in]) 1);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    37
qed_spec_mp "merges_incr";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    38
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    39
(* is redundant but useful *)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    40
Goal "[| xs!i = Some x; xs : listsn n (option A); i < n |] ==> x:A";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    41
by (dtac listsnE_nth_in 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    42
by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    43
by (Asm_full_simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    44
qed "listsn_optionE_in";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    45
(*Addsimps [listsn_optionE_in];*)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    46
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    47
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    48
 "[| semilat L |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    49
\ (!xs. xs : listsn n (option L) --> x:L --> (!p:set ps. p<n) --> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    50
\      (merges x ps xs = xs) = (!p:set ps. Some x <= xs!p))";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    51
by (induct_tac "ps" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    52
 by (Asm_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    53
by (Clarsimp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    54
by (rename_tac "p ps xs" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    55
by (rtac iffI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    56
 by (rtac context_conjI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    57
  by (subgoal_tac "xs[p := Some x + xs!p] <= xs" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    58
   by (EVERY[etac subst 2, rtac merges_incr 2]);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    59
      by (assume_tac 2);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    60
     by (Force_tac 2);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    61
    by (assume_tac 2);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    62
   by (assume_tac 2);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    63
  by (exhaust_tac "xs!p" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    64
   by (asm_full_simp_tac (simpset() addsimps [list_update_le_conv]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    65
  by (asm_full_simp_tac (simpset() addsimps
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    66
      [list_update_le_conv,listsn_optionE_in]) 1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    67
 by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    68
 by (rotate_tac ~3 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    69
 by (asm_full_simp_tac (simpset() addsimps
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    70
      [le_iff_plus_unchanged RS iffD1,listsn_optionE_in,
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    71
       list_update_same_conv RS iffD2]) 1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    72
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    73
by (asm_simp_tac (simpset() addsimps
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    74
      [le_iff_plus_unchanged RS iffD1,listsn_optionE_in,
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    75
       list_update_same_conv RS iffD2]) 1);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    76
qed_spec_mp "merges_same_conv";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    77
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    78
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    79
Goalw [le_list,plus_option]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    80
 "xs <= ys --> xs : listsn n (option L) --> ys : listsn n (option L) --> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    81
\ p < n --> ys!p = Some y --> x <= y --> x : L --> semilat L --> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    82
\ xs[p := Some x + xs!p] <= ys";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    83
by (simp_tac (simpset()  addsimps [nth_list_update]
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    84
                        addsplits [option.split]) 1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    85
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    86
by (rotate_tac 3 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    87
by (force_tac (claset() addEs [listsn_optionE_in] addIs [semilat_lub],
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    88
              simpset()) 1);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    89
qed_spec_mp "list_update_le_listI";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    90
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    91
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    92
 "[| semilat(L); t:L; tos : listsn n (option L); \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    93
\    !p. p:set ps --> Some t <= tos!p; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    94
\    !p. p:set ps --> p<n |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    95
\ set qs <= set ps  --> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    96
\ (!sos. sos : listsn n (option L) & sos <= tos --> merges t qs sos <= tos)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    97
by (induct_tac "qs" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    98
 by (Asm_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    99
by (force_tac (claset(),simpset() addsimps [list_update_le_listI]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   100
val lemma = result();
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   101
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   102
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   103
 "[| semilat(L); t:L; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   104
\    !p. p:set ps --> Some t <= tos!p; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   105
\    !p. p:set ps --> p<n; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   106
\    sos <= tos; sos : listsn n (option L); tos : listsn n (option L) |] \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   107
\ ==> merges t ps sos <= tos";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   108
by (blast_tac (claset() addDs [lemma]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   109
qed "merges_pres_le_ub";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   110
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   111
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   112
(** next **)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   113
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   114
Goalw [is_next_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   115
 "[| is_next next; next step succs sos = None; succs_bounded succs n; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   116
\    sos : listsn n S |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   117
\ ? p<n. ? s. sos!p = Some s & step p s = None";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   118
by (subgoal_tac "n=size sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   119
by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   120
by (Asm_simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   121
qed "next_None";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   122
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   123
Goalw [is_next_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   124
 "[| is_next next; succs_bounded succs n; sos : listsn n S |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   125
\ next step succs sos = Some sos --> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   126
\ (!p<n. !s. sos!p = Some s --> (? t. \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   127
\         step p s = Some(t) & merges t (succs p) sos = sos))";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   128
by (subgoal_tac "n=size sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   129
by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   130
by (Asm_simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   131
qed "next_Some1";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   132
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   133
Goalw [is_next_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   134
 "[| is_next next; next step succs sos = Some sos'; sos' ~= sos; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   135
\    succs_bounded succs n; sos : listsn n S |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   136
\ ? p<n. ? s. sos!p = Some s & (? t. \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   137
\     step p s = Some(t) & merges t (succs p) sos = sos')";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   138
by (subgoal_tac "n=size sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   139
by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   140
by (Asm_simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   141
qed "next_Some2";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   142
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   143
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   144
 "[| is_next next; succs_bounded succs n; sos : listsn n S |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   145
\ (next step succs sos = Some sos) = \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   146
\ (!p<n. !s. sos!p = Some s --> (? t. \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   147
\         step p s = Some(t) & merges t (succs p) sos = sos))";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   148
by (rtac iffI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   149
 by (asm_simp_tac (simpset() addsimps [next_Some1]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   150
by (exhaust_tac "next step succs sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   151
 by (best_tac (claset() addSDs [next_None] addss simpset()) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   152
by (rename_tac "sos'" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   153
by (case_tac "sos' = sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   154
 by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   155
by (best_tac (claset() addSDs [next_Some2] addss simpset()) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   156
qed "next_Some1_eq";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   157
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   158
Addsimps [next_Some1_eq];
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   159
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   160
Goalw [step_pres_type_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   161
 "[| step_pres_type step n L; s:L; p<n; step p s = Some(t) |] ==> t:L";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   162
by (Blast_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   163
qed "step_pres_typeD";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   164
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   165
Goalw [succs_bounded_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   166
 "[| succs_bounded succs n; p < n; q : set(succs p) |] ==> q < n";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   167
by (Blast_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   168
qed "succs_boundedD";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   169
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   170
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   171
 "[| is_next next; semilat A; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   172
\    step_pres_type step n A; succs_bounded succs n;\
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   173
\    sos : listsn n (option A) |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   174
\ next step succs sos : option (listsn n (option A))";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   175
by (exhaust_tac "next step succs sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   176
 by (ALLGOALS Asm_simp_tac);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   177
by (rename_tac "sos'" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   178
by (case_tac "sos' = sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   179
 by (Asm_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   180
by (blast_tac (claset() addIs [step_pres_typeD,succs_boundedD,listsn_optionE_in] addSIs [merges_preserves_type] addDs [next_Some2]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   181
qed_spec_mp "next_preserves_type";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   182
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   183
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   184
 "[| is_next next; semilat A; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   185
\    step_pres_type step n A; succs_bounded succs n; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   186
\    next step succs xs = Some ys; xs : listsn n (option A) |] ==> xs <= ys";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   187
by (case_tac "ys = xs" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   188
 by (Asm_full_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   189
by (blast_tac (claset() addSIs [merges_incr] addIs [listsn_optionE_in]
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   190
       addDs [step_pres_typeD,succs_boundedD,next_Some2]) 1);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   191
qed_spec_mp "next_incr";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   192
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   193
val lemma = (Unify.trace_bound, Unify.search_bound);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   194
Unify.trace_bound := 50;
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   195
Unify.search_bound := 50;
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   196
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   197
Goalw [is_next_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   198
 "is_next (%step succs sos. itnext (size sos) step succs sos)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   199
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   200
by (etac thin_rl 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   201
by (res_inst_tac [("n","length sos")] nat_induct 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   202
 by (Asm_full_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   203
by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq,Let_def]
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   204
                                addsplits [option.split])1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   205
by (Blast_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   206
qed "is_next_itnext";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   207
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   208
Unify.trace_bound := !(fst lemma);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   209
Unify.search_bound := !(snd lemma);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   210
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   211
(** fix step **)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   212
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   213
Goalw [step_mono_None_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   214
 "[| step_mono_None step n L; s : L; p < n; s <= t; step p s = None |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   215
\ step p t = None";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   216
by (Blast_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   217
qed "step_mono_NoneD";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   218
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   219
Goalw [step_mono_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   220
 "[| step_mono step n L; s : L; p < n; s <= t; step p s = Some(u) |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   221
\ !v. step p t = Some(v) --> u <= v";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   222
by (Blast_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   223
qed "step_monoD";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   224
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   225
Goalw [stable_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   226
"[| is_next next; semilat L; sos : listsn n (option L); \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   227
\   step_pres_type step n L; succs_bounded succs n |] \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   228
\ ==> (next step succs sos = Some sos) = (!p<n. stable step succs p sos)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   229
by (Asm_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   230
by (rtac iffI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   231
 by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   232
 by (etac allE 1 THEN mp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   233
 by (etac allE 1 THEN mp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   234
 by (Clarify_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   235
 bd(merges_same_conv RS iffD1)1;
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   236
     by (assume_tac 4);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   237
    by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   238
   by (blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   239
  by (blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   240
 by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   241
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   242
by (etac allE 1 THEN mp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   243
by (Asm_full_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   244
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   245
by (Asm_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   246
by (blast_tac (claset() addSIs [merges_same_conv RS iffD2]
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   247
     addIs [step_pres_typeD,succs_boundedD,listsn_optionE_in]) 1);
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   248
qed "fixpoint_next_iff_stable";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   249
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   250
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   251
 "[| semilat L; is_next next; succs_bounded succs n; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   252
\    step_pres_type step n L; step_mono_None step n L; step_mono step n L; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   253
\    tos:listsn n (option L); next step succs tos = Some tos; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   254
\    sos:listsn n (option L); sos <= tos |] \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   255
\ ==> ? sos'. next step succs sos = Some sos' & sos' <= tos";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   256
by (subgoal_tac
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   257
   "!p<n. !s. sos!p = Some s --> (? u. \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   258
\             step p s = Some(u) & (!q:set(succs p). Some u<=tos!q))" 1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   259
 by (exhaust_tac "next step succs sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   260
  by (dtac next_None 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   261
     by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   262
    by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   263
   by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   264
  by (Force_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   265
 by (rename_tac "sos'" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   266
 by (case_tac "sos' = sos" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   267
  by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   268
 by (dtac next_Some2 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   269
    by (EVERY1[atac, atac, atac, atac]);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   270
 by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   271
 by (etac allE 1 THEN mp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   272
 by (etac allE 1 THEN mp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   273
 by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   274
 by (EVERY1[rtac exI, rtac conjI, atac]);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   275
 by (rtac merges_pres_le_ub 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   276
       by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   277
      by (blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   278
     by (Asm_full_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   279
    by (blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   280
   by (REPEAT(atac 1));
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   281
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   282
by (exhaust_tac "tos!p" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   283
 by (force_tac (claset() addDs [le_listD],simpset()) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   284
by (rename_tac "t" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   285
by (subgoal_tac "s <= t" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   286
 by (force_tac (claset() addDs [le_listD],simpset()) 2);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   287
by (exhaust_tac "step p s" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   288
 by (dtac step_mono_NoneD 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   289
     by (assume_tac 4);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   290
    by (blast_tac (claset() addIs [listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   291
   by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   292
  by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   293
 by (Force_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   294
by (dtac step_monoD 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   295
    by (assume_tac 4);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   296
   by (blast_tac (claset() addIs [listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   297
  by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   298
 by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   299
by (Asm_full_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   300
by (etac allE 1 THEN mp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   301
by (etac allE 1 THEN mp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   302
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   303
by (Asm_full_simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   304
by (forward_tac[merges_same_conv RS iffD1]1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   305
    by (assume_tac 4);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   306
   by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   307
  by (blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   308
 by (blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   309
by (blast_tac (claset() addIs [order_trans]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   310
val lemma = result();
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   311
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   312
Goalw [is_dfa_def]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   313
 "[| semilat L; acc L; is_next next; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   314
\    step_pres_type step n L; succs_bounded succs n; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   315
\    step_mono_None step n L; step_mono step n L |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   316
\ is_dfa (%sos. fix(next step succs, sos)) step succs n L";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   317
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   318
by (stac fix_iff_has_fixpoint 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   319
     by (etac (acc_option RS acc_listsn) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   320
    by (blast_tac (claset() addIs [lemma]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   321
   by (blast_tac (claset() addIs [next_preserves_type]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   322
  by (blast_tac (claset() addIs [next_incr]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   323
 by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   324
by (asm_simp_tac (simpset() delsimps [next_Some1_eq] addsimps [fixpoint_next_iff_stable]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   325
qed "is_dfa_fix_next";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   326
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   327
Goal
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   328
 "[| semilat L; acc L; is_next next; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   329
\    step_pres_type step n L; succs_bounded succs n; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   330
\    step_mono_None step n L; step_mono step n L; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   331
\    wti_is_fix_step step wti succs n L; \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   332
\    sos : listsn n (option L) |] ==> \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   333
\ fix(next step succs, sos) = \
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   334
\ (? tos:listsn n (option L). sos<=tos & welltyping wti tos)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   335
by (blast_tac (claset() addSDs [dfa_iff_welltyping] addSIs [is_dfa_fix_next]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   336
qed "fix_next_iff_welltyping";