thys/LexerExt.thy
author cu
Sun, 08 Oct 2017 14:21:24 +0100
changeset 277 42268a284ea6
parent 276 a3134f7de065
child 278 424bdcd01016
permissions -rw-r--r--
updated
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
   
220
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
     2
theory LexerExt
276
a3134f7de065 updated
cu
parents: 261
diff changeset
     3
  imports SpecExt 
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
begin
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
276
a3134f7de065 updated
cu
parents: 261
diff changeset
     7
section {* The Lexer Functions by Sulzmann and Lu  *}
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
fun 
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
  mkeps :: "rexp \<Rightarrow> val"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
where
107
6adda4a667b1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
    12
  "mkeps(ONE) = Void"
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    14
| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
97
38696f516c6b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 95
diff changeset
    15
| "mkeps(STAR r) = Stars []"
220
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
    16
| "mkeps(UPNTIMES r n) = Stars []"
221
c21938d0b396 added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents: 220
diff changeset
    17
| "mkeps(NTIMES r n) = Stars (replicate n (mkeps r))"
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
    18
| "mkeps(FROMNTIMES r n) = Stars (replicate n (mkeps r))"
227
a8749366ad5d NMTIMES case also done
Christian Urban <urbanc@in.tum.de>
parents: 226
diff changeset
    19
| "mkeps(NMTIMES r n m) = Stars (replicate n (mkeps r))"
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    20
  
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
where
101
7f4f8c34da95 fixed inj function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
    23
  "injval (CHAR d) c Void = Char d"
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
| "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
| "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
| "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
| "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
| "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
97
38696f516c6b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 95
diff changeset
    29
| "injval (STAR r) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" 
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    30
| "injval (NTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
    31
| "injval (FROMNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
    32
| "injval (UPNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
    33
| "injval (NMTIMES r n m) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
    34
  
a3134f7de065 updated
cu
parents: 261
diff changeset
    35
fun 
a3134f7de065 updated
cu
parents: 261
diff changeset
    36
  lexer :: "rexp \<Rightarrow> string \<Rightarrow> val option"
a3134f7de065 updated
cu
parents: 261
diff changeset
    37
where
a3134f7de065 updated
cu
parents: 261
diff changeset
    38
  "lexer r [] = (if nullable r then Some(mkeps r) else None)"
a3134f7de065 updated
cu
parents: 261
diff changeset
    39
| "lexer r (c#s) = (case (lexer (der c r) s) of  
a3134f7de065 updated
cu
parents: 261
diff changeset
    40
                    None \<Rightarrow> None
a3134f7de065 updated
cu
parents: 261
diff changeset
    41
                  | Some(v) \<Rightarrow> Some(injval r c v))"
107
6adda4a667b1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
    42
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    43
a3134f7de065 updated
cu
parents: 261
diff changeset
    44
a3134f7de065 updated
cu
parents: 261
diff changeset
    45
section {* Mkeps, Injval Properties *}
243
09ab631ce7fa updated literature
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
    46
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
lemma mkeps_flat:
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
  assumes "nullable(r)" 
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
  shows "flat (mkeps r) = []"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
using assms
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    51
  apply(induct rule: nullable.induct) 
a3134f7de065 updated
cu
parents: 261
diff changeset
    52
         apply(auto)
a3134f7de065 updated
cu
parents: 261
diff changeset
    53
  by presburger  
a3134f7de065 updated
cu
parents: 261
diff changeset
    54
  
a3134f7de065 updated
cu
parents: 261
diff changeset
    55
  
a3134f7de065 updated
cu
parents: 261
diff changeset
    56
lemma mkeps_nullable:
a3134f7de065 updated
cu
parents: 261
diff changeset
    57
  assumes "nullable(r)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
    58
  shows "\<Turnstile> mkeps r : r"
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
using assms
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    60
apply(induct rule: nullable.induct) 
a3134f7de065 updated
cu
parents: 261
diff changeset
    61
         apply(auto intro: Prf.intros split: if_splits)
a3134f7de065 updated
cu
parents: 261
diff changeset
    62
  using Prf.intros(8) apply force
a3134f7de065 updated
cu
parents: 261
diff changeset
    63
     apply(subst append.simps(1)[symmetric])
a3134f7de065 updated
cu
parents: 261
diff changeset
    64
     apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
    65
       apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    66
      apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    67
       apply (simp add: mkeps_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
    68
      apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    69
  using Prf.intros(9) apply force
a3134f7de065 updated
cu
parents: 261
diff changeset
    70
    apply(subst append.simps(1)[symmetric])
a3134f7de065 updated
cu
parents: 261
diff changeset
    71
     apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
    72
       apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    73
      apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    74
       apply (simp add: mkeps_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
    75
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    76
  using Prf.intros(11) apply force
a3134f7de065 updated
cu
parents: 261
diff changeset
    77
    apply(subst append.simps(1)[symmetric])
a3134f7de065 updated
cu
parents: 261
diff changeset
    78
     apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
    79
       apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    80
      apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    81
    apply (simp add: mkeps_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
    82
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
    83
  apply(simp)
227
a8749366ad5d NMTIMES case also done
Christian Urban <urbanc@in.tum.de>
parents: 226
diff changeset
    84
done
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    85
    
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
243
09ab631ce7fa updated literature
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
    87
lemma Prf_injval_flat:
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    88
  assumes "\<Turnstile> v : der c r" 
243
09ab631ce7fa updated literature
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
    89
  shows "flat (injval r c v) = c # (flat v)"
09ab631ce7fa updated literature
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
    90
using assms
09ab631ce7fa updated literature
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
    91
apply(induct arbitrary: v rule: der.induct)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
    92
apply(auto elim!: Prf_elims intro: mkeps_flat split: if_splits)
a3134f7de065 updated
cu
parents: 261
diff changeset
    93
done
a3134f7de065 updated
cu
parents: 261
diff changeset
    94
a3134f7de065 updated
cu
parents: 261
diff changeset
    95
lemma Prf_injval:
a3134f7de065 updated
cu
parents: 261
diff changeset
    96
  assumes "\<Turnstile> v : der c r" 
a3134f7de065 updated
cu
parents: 261
diff changeset
    97
  shows "\<Turnstile> (injval r c v) : r"
a3134f7de065 updated
cu
parents: 261
diff changeset
    98
using assms
a3134f7de065 updated
cu
parents: 261
diff changeset
    99
apply(induct r arbitrary: c v rule: rexp.induct)
a3134f7de065 updated
cu
parents: 261
diff changeset
   100
apply(auto intro!: Prf.intros mkeps_nullable elim!: Prf_elims split: if_splits)[6]
a3134f7de065 updated
cu
parents: 261
diff changeset
   101
    apply(simp add: Prf_injval_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
   102
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   103
  apply(case_tac x2)
a3134f7de065 updated
cu
parents: 261
diff changeset
   104
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   105
  apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   106
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   107
   apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   108
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   109
  apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   110
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   111
  using Prf.intros(7) Prf_injval_flat apply auto[1]
a3134f7de065 updated
cu
parents: 261
diff changeset
   112
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   113
    apply(case_tac x2)
a3134f7de065 updated
cu
parents: 261
diff changeset
   114
     apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   115
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   116
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   117
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   118
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   119
  apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   120
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   121
    apply(subst append.simps(2)[symmetric])
a3134f7de065 updated
cu
parents: 261
diff changeset
   122
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
   123
      apply(simp add: Prf_injval_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
   124
     apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   125
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   126
    prefer 2
a3134f7de065 updated
cu
parents: 261
diff changeset
   127
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   128
   apply(case_tac "x3a < x2")
a3134f7de065 updated
cu
parents: 261
diff changeset
   129
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   130
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   131
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   132
    apply(case_tac x2)
a3134f7de065 updated
cu
parents: 261
diff changeset
   133
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   134
    apply(case_tac x3a)
a3134f7de065 updated
cu
parents: 261
diff changeset
   135
     apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   136
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   137
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   138
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   139
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   140
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   141
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   142
  using Prf.intros(12) Prf_injval_flat apply auto[1]
a3134f7de065 updated
cu
parents: 261
diff changeset
   143
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   144
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   145
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   146
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   147
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   148
    apply(subst append.simps(2)[symmetric])
a3134f7de065 updated
cu
parents: 261
diff changeset
   149
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
   150
     apply(simp add: Prf_injval_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
   151
     apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   152
     apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   153
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   154
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   155
  using Prf.intros(12) Prf_injval_flat apply auto[1]
a3134f7de065 updated
cu
parents: 261
diff changeset
   156
    apply(case_tac x2)
a3134f7de065 updated
cu
parents: 261
diff changeset
   157
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   158
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   159
   apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   160
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   161
    apply(simp_all)
a3134f7de065 updated
cu
parents: 261
diff changeset
   162
    apply (simp add: Prf.intros(10) Prf_injval_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
   163
  using Prf.intros(10) Prf_injval_flat apply auto[1]
a3134f7de065 updated
cu
parents: 261
diff changeset
   164
  apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   165
  apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   166
    apply(erule Prf_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   167
    apply(simp_all)
a3134f7de065 updated
cu
parents: 261
diff changeset
   168
    apply(subst append.simps(2)[symmetric])
a3134f7de065 updated
cu
parents: 261
diff changeset
   169
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
   170
     apply(simp add: Prf_injval_flat)
a3134f7de065 updated
cu
parents: 261
diff changeset
   171
     apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   172
   apply(simp)
243
09ab631ce7fa updated literature
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   173
done
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   177
text {*
a3134f7de065 updated
cu
parents: 261
diff changeset
   178
  Mkeps and injval produce, or preserve, Posix values.
a3134f7de065 updated
cu
parents: 261
diff changeset
   179
*}
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   180
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   181
lemma Posix_mkeps:
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
  assumes "nullable r"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
  shows "[] \<in> r \<rightarrow> mkeps r"
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
using assms
185
841f7b9c0a6a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 172
diff changeset
   185
apply(induct r rule: nullable.induct)
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   186
apply(auto intro: Posix.intros simp add: nullable_correctness Sequ_def)
89
9613e6ace30f added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
apply(subst append.simps(1)[symmetric])
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   188
apply(rule Posix.intros)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   189
      apply(auto)
a3134f7de065 updated
cu
parents: 261
diff changeset
   190
  done
a3134f7de065 updated
cu
parents: 261
diff changeset
   191
    
100
8b919b3d753e strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 99
diff changeset
   192
172
cdc0bdcfba3f updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   193
lemma Posix_injval:
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   194
  assumes "s \<in> (der c r) \<rightarrow> v" 
143
1e7b36450d9a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 142
diff changeset
   195
  shows "(c # s) \<in> r \<rightarrow> (injval r c v)"
106
489dfa0d7ec9 more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
   196
using assms
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   197
proof(induct r arbitrary: s v rule: rexp.induct)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   198
  case ZERO
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   199
  have "s \<in> der c ZERO \<rightarrow> v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   200
  then have "s \<in> ZERO \<rightarrow> v" by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   201
  then have "False" by cases
143
1e7b36450d9a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 142
diff changeset
   202
  then show "(c # s) \<in> ZERO \<rightarrow> (injval ZERO c v)" by simp
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   203
next
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   204
  case ONE
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   205
  have "s \<in> der c ONE \<rightarrow> v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   206
  then have "s \<in> ZERO \<rightarrow> v" by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   207
  then have "False" by cases
143
1e7b36450d9a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 142
diff changeset
   208
  then show "(c # s) \<in> ONE \<rightarrow> (injval ONE c v)" by simp
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   209
next 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   210
  case (CHAR d)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   211
  consider (eq) "c = d" | (ineq) "c \<noteq> d" by blast
143
1e7b36450d9a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 142
diff changeset
   212
  then show "(c # s) \<in> (CHAR d) \<rightarrow> (injval (CHAR d) c v)"
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   213
  proof (cases)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   214
    case eq
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   215
    have "s \<in> der c (CHAR d) \<rightarrow> v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   216
    then have "s \<in> ONE \<rightarrow> v" using eq by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   217
    then have eqs: "s = [] \<and> v = Void" by cases simp
142
08dcf0d20f15 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 127
diff changeset
   218
    show "(c # s) \<in> CHAR d \<rightarrow> injval (CHAR d) c v" using eq eqs 
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   219
    by (auto intro: Posix.intros)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   220
  next
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   221
    case ineq
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   222
    have "s \<in> der c (CHAR d) \<rightarrow> v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   223
    then have "s \<in> ZERO \<rightarrow> v" using ineq by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   224
    then have "False" by cases
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   225
    then show "(c # s) \<in> CHAR d \<rightarrow> injval (CHAR d) c v" by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   226
  qed
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   227
next
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   228
  case (ALT r1 r2)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   229
  have IH1: "\<And>s v. s \<in> der c r1 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r1 \<rightarrow> injval r1 c v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   230
  have IH2: "\<And>s v. s \<in> der c r2 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r2 \<rightarrow> injval r2 c v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   231
  have "s \<in> der c (ALT r1 r2) \<rightarrow> v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   232
  then have "s \<in> ALT (der c r1) (der c r2) \<rightarrow> v" by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   233
  then consider (left) v' where "v = Left v'" "s \<in> der c r1 \<rightarrow> v'" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   234
              | (right) v' where "v = Right v'" "s \<notin> L (der c r1)" "s \<in> der c r2 \<rightarrow> v'" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   235
              by cases auto
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   236
  then show "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c v"
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   237
  proof (cases)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   238
    case left
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   239
    have "s \<in> der c r1 \<rightarrow> v'" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   240
    then have "(c # s) \<in> r1 \<rightarrow> injval r1 c v'" using IH1 by simp
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   241
    then have "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c (Left v')" by (auto intro: Posix.intros)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   242
    then show "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c v" using left by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   243
  next 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   244
    case right
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   245
    have "s \<notin> L (der c r1)" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   246
    then have "c # s \<notin> L r1" by (simp add: der_correctness Der_def)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   247
    moreover 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   248
    have "s \<in> der c r2 \<rightarrow> v'" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   249
    then have "(c # s) \<in> r2 \<rightarrow> injval r2 c v'" using IH2 by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   250
    ultimately have "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c (Right v')" 
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   251
      by (auto intro: Posix.intros)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   252
    then show "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c v" using right by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   253
  qed
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   254
next
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   255
  case (SEQ r1 r2)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   256
  have IH1: "\<And>s v. s \<in> der c r1 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r1 \<rightarrow> injval r1 c v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   257
  have IH2: "\<And>s v. s \<in> der c r2 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r2 \<rightarrow> injval r2 c v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   258
  have "s \<in> der c (SEQ r1 r2) \<rightarrow> v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   259
  then consider 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   260
        (left_nullable) v1 v2 s1 s2 where 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   261
        "v = Left (Seq v1 v2)"  "s = s1 @ s2" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   262
        "s1 \<in> der c r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" "nullable r1" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   263
        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r1) \<and> s\<^sub>4 \<in> L r2)"
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   264
      | (right_nullable) v1 s1 s2 where 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   265
        "v = Right v1" "s = s1 @ s2"  
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   266
        "s \<in> der c r2 \<rightarrow> v1" "nullable r1" "s1 @ s2 \<notin> L (SEQ (der c r1) r2)"
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   267
      | (not_nullable) v1 v2 s1 s2 where
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   268
        "v = Seq v1 v2" "s = s1 @ s2" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   269
        "s1 \<in> der c r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" "\<not>nullable r1" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   270
        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r1) \<and> s\<^sub>4 \<in> L r2)"
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   271
        by (force split: if_splits elim!: Posix_elims simp add: Sequ_def der_correctness Der_def)   
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   272
  then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   273
    proof (cases)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   274
      case left_nullable
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   275
      have "s1 \<in> der c r1 \<rightarrow> v1" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   276
      then have "(c # s1) \<in> r1 \<rightarrow> injval r1 c v1" using IH1 by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   277
      moreover
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   278
      have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r1) \<and> s\<^sub>4 \<in> L r2)" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   279
      then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by (simp add: der_correctness Der_def)
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   280
      ultimately have "((c # s1) @ s2) \<in> SEQ r1 r2 \<rightarrow> Seq (injval r1 c v1) v2" using left_nullable by (rule_tac Posix.intros)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   281
      then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" using left_nullable by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   282
    next
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   283
      case right_nullable
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   284
      have "nullable r1" by fact
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   285
      then have "[] \<in> r1 \<rightarrow> (mkeps r1)" by (rule Posix_mkeps)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   286
      moreover
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   287
      have "s \<in> der c r2 \<rightarrow> v1" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   288
      then have "(c # s) \<in> r2 \<rightarrow> (injval r2 c v1)" using IH2 by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   289
      moreover
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   290
      have "s1 @ s2 \<notin> L (SEQ (der c r1) r2)" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   291
      then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = c # s \<and> [] @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" using right_nullable
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   292
        by(auto simp add: der_correctness Der_def append_eq_Cons_conv Sequ_def)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   293
      ultimately have "([] @ (c # s)) \<in> SEQ r1 r2 \<rightarrow> Seq (mkeps r1) (injval r2 c v1)"
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   294
      by(rule Posix.intros)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   295
      then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" using right_nullable by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   296
    next
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   297
      case not_nullable
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   298
      have "s1 \<in> der c r1 \<rightarrow> v1" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   299
      then have "(c # s1) \<in> r1 \<rightarrow> injval r1 c v1" using IH1 by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   300
      moreover
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   301
      have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r1) \<and> s\<^sub>4 \<in> L r2)" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   302
      then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by (simp add: der_correctness Der_def)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   303
      ultimately have "((c # s1) @ s2) \<in> SEQ r1 r2 \<rightarrow> Seq (injval r1 c v1) v2" using not_nullable 
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   304
        by (rule_tac Posix.intros) (simp_all) 
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   305
      then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" using not_nullable by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   306
    qed
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   307
next
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   308
case (UPNTIMES r n s v)
a3134f7de065 updated
cu
parents: 261
diff changeset
   309
  have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   310
  have "s \<in> der c (UPNTIMES r n) \<rightarrow> v" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   311
  then consider
a3134f7de065 updated
cu
parents: 261
diff changeset
   312
      (cons) v1 vs s1 s2 where 
a3134f7de065 updated
cu
parents: 261
diff changeset
   313
        "v = Seq v1 (Stars vs)" "s = s1 @ s2" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   314
        "s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (UPNTIMES r (n - 1)) \<rightarrow> (Stars vs)" "0 < n"
a3134f7de065 updated
cu
parents: 261
diff changeset
   315
        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (UPNTIMES r (n - 1)))" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   316
    (* here *)
a3134f7de065 updated
cu
parents: 261
diff changeset
   317
    apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits)
a3134f7de065 updated
cu
parents: 261
diff changeset
   318
    apply(erule Posix_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   319
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   320
    apply(subgoal_tac "\<exists>vss. v2 = Stars vss")
a3134f7de065 updated
cu
parents: 261
diff changeset
   321
    apply(clarify)
a3134f7de065 updated
cu
parents: 261
diff changeset
   322
    apply(drule_tac x="v1" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   323
    apply(drule_tac x="vss" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   324
    apply(drule_tac x="s1" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   325
    apply(drule_tac x="s2" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   326
     apply(simp add: der_correctness Der_def)
a3134f7de065 updated
cu
parents: 261
diff changeset
   327
    apply(erule Posix_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   328
     apply(auto)
a3134f7de065 updated
cu
parents: 261
diff changeset
   329
      done
a3134f7de065 updated
cu
parents: 261
diff changeset
   330
    then show "(c # s) \<in> (UPNTIMES r n) \<rightarrow> injval (UPNTIMES r n) c v" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   331
    proof (cases)
a3134f7de065 updated
cu
parents: 261
diff changeset
   332
      case cons
a3134f7de065 updated
cu
parents: 261
diff changeset
   333
          have "s1 \<in> der c r \<rightarrow> v1" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   334
          then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp
a3134f7de065 updated
cu
parents: 261
diff changeset
   335
        moreover
a3134f7de065 updated
cu
parents: 261
diff changeset
   336
          have "s2 \<in> (UPNTIMES r (n - 1)) \<rightarrow> Stars vs" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   337
        moreover 
a3134f7de065 updated
cu
parents: 261
diff changeset
   338
          have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact 
a3134f7de065 updated
cu
parents: 261
diff changeset
   339
          then have "flat (injval r c v1) = (c # s1)" by (rule Posix1)
a3134f7de065 updated
cu
parents: 261
diff changeset
   340
          then have "flat (injval r c v1) \<noteq> []" by simp
a3134f7de065 updated
cu
parents: 261
diff changeset
   341
        moreover 
a3134f7de065 updated
cu
parents: 261
diff changeset
   342
          have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (UPNTIMES r (n - 1)))" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   343
          then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (n - 1)))" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   344
            by (simp add: der_correctness Der_def)
a3134f7de065 updated
cu
parents: 261
diff changeset
   345
        ultimately 
a3134f7de065 updated
cu
parents: 261
diff changeset
   346
        have "((c # s1) @ s2) \<in> UPNTIMES r n \<rightarrow> Stars (injval r c v1 # vs)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   347
           thm Posix.intros
a3134f7de065 updated
cu
parents: 261
diff changeset
   348
           apply (rule_tac Posix.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
   349
               apply(simp_all)
a3134f7de065 updated
cu
parents: 261
diff changeset
   350
              apply(case_tac n)
a3134f7de065 updated
cu
parents: 261
diff changeset
   351
            apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   352
           using Posix_elims(1) UPNTIMES.prems apply auto[1]
a3134f7de065 updated
cu
parents: 261
diff changeset
   353
             apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   354
             done
a3134f7de065 updated
cu
parents: 261
diff changeset
   355
        then show "(c # s) \<in> UPNTIMES r n \<rightarrow> injval (UPNTIMES r n) c v" using cons by(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   356
      qed
a3134f7de065 updated
cu
parents: 261
diff changeset
   357
    next
a3134f7de065 updated
cu
parents: 261
diff changeset
   358
      case (STAR r)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   359
  have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   360
  have "s \<in> der c (STAR r) \<rightarrow> v" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   361
  then consider
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   362
      (cons) v1 vs s1 s2 where 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   363
        "v = Seq v1 (Stars vs)" "s = s1 @ s2" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   364
        "s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (STAR r) \<rightarrow> (Stars vs)"
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   365
        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (STAR r))" 
149
ec3d221bfc45 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 146
diff changeset
   366
        apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros)
142
08dcf0d20f15 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 127
diff changeset
   367
        apply(rotate_tac 3)
149
ec3d221bfc45 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 146
diff changeset
   368
        apply(erule_tac Posix_elims(6))
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   369
        apply (simp add: Posix.intros(6))
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   370
        using Posix.intros(7) by blast
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   371
    then show "(c # s) \<in> STAR r \<rightarrow> injval (STAR r) c v" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   372
    proof (cases)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   373
      case cons
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   374
          have "s1 \<in> der c r \<rightarrow> v1" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   375
          then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   376
        moreover
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   377
          have "s2 \<in> STAR r \<rightarrow> Stars vs" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   378
        moreover 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   379
          have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact 
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   380
          then have "flat (injval r c v1) = (c # s1)" by (rule Posix1)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   381
          then have "flat (injval r c v1) \<noteq> []" by simp
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   382
        moreover 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   383
          have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (STAR r))" by fact
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   384
          then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" 
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   385
            by (simp add: der_correctness Der_def)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   386
        ultimately 
146
da81ffac4b10 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
   387
        have "((c # s1) @ s2) \<in> STAR r \<rightarrow> Stars (injval r c v1 # vs)" by (rule Posix.intros)
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   388
        then show "(c # s) \<in> STAR r \<rightarrow> injval (STAR r) c v" using cons by(simp)
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   389
    qed
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   390
  next
a3134f7de065 updated
cu
parents: 261
diff changeset
   391
    case (NTIMES r n s v)
a3134f7de065 updated
cu
parents: 261
diff changeset
   392
     have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   393
  have "s \<in> der c (NTIMES r n) \<rightarrow> v" by fact
220
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   394
  then consider
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   395
      (cons) v1 vs s1 s2 where 
220
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   396
        "v = Seq v1 (Stars vs)" "s = s1 @ s2" 
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   397
        "s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (NTIMES r (n - 1)) \<rightarrow> (Stars vs)" "0 < n"
a3134f7de065 updated
cu
parents: 261
diff changeset
   398
        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   399
    apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits)
a3134f7de065 updated
cu
parents: 261
diff changeset
   400
    apply(erule Posix_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   401
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   402
    apply(subgoal_tac "\<exists>vss. v2 = Stars vss")
a3134f7de065 updated
cu
parents: 261
diff changeset
   403
    apply(clarify)
a3134f7de065 updated
cu
parents: 261
diff changeset
   404
    apply(drule_tac x="v1" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   405
    apply(drule_tac x="vss" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   406
    apply(drule_tac x="s1" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   407
    apply(drule_tac x="s2" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   408
     apply(simp add: der_correctness Der_def)
a3134f7de065 updated
cu
parents: 261
diff changeset
   409
    apply(erule Posix_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   410
     apply(auto)
a3134f7de065 updated
cu
parents: 261
diff changeset
   411
      done
a3134f7de065 updated
cu
parents: 261
diff changeset
   412
    then show "(c # s) \<in> (NTIMES r n) \<rightarrow> injval (NTIMES r n) c v" 
220
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   413
    proof (cases)
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   414
      case cons
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   415
          have "s1 \<in> der c r \<rightarrow> v1" by fact
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   416
          then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   417
        moreover
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   418
          have "s2 \<in> (NTIMES r (n - 1)) \<rightarrow> Stars vs" by fact
220
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   419
        moreover 
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   420
          have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact 
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   421
          then have "flat (injval r c v1) = (c # s1)" by (rule Posix1)
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   422
          then have "flat (injval r c v1) \<noteq> []" by simp
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   423
        moreover 
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   424
          have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   425
          then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))" 
220
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   426
            by (simp add: der_correctness Der_def)
a8b32da484df updated
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
   427
        ultimately 
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   428
        have "((c # s1) @ s2) \<in> NTIMES r n \<rightarrow> Stars (injval r c v1 # vs)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   429
           apply (rule_tac Posix.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
   430
               apply(simp_all)
a3134f7de065 updated
cu
parents: 261
diff changeset
   431
              apply(case_tac n)
a3134f7de065 updated
cu
parents: 261
diff changeset
   432
            apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   433
           using Posix_elims(1) NTIMES.prems apply auto[1]
a3134f7de065 updated
cu
parents: 261
diff changeset
   434
             apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   435
             done
221
c21938d0b396 added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents: 220
diff changeset
   436
        then show "(c # s) \<in> NTIMES r n \<rightarrow> injval (NTIMES r n) c v" using cons by(simp)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   437
      qed  
a3134f7de065 updated
cu
parents: 261
diff changeset
   438
  next
a3134f7de065 updated
cu
parents: 261
diff changeset
   439
    case (FROMNTIMES r n s v)
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   440
  have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   441
  have "s \<in> der c (FROMNTIMES r n) \<rightarrow> v" by fact
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   442
  then consider
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   443
      (cons) v1 vs s1 s2 where 
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   444
        "v = Seq v1 (Stars vs)" "s = s1 @ s2" 
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   445
        "s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (FROMNTIMES r (n - 1)) \<rightarrow> (Stars vs)" "0 < n"
a3134f7de065 updated
cu
parents: 261
diff changeset
   446
        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (FROMNTIMES r (n - 1)))"
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   447
     | (null) v1 vs s1 s2 where 
42268a284ea6 updated
cu
parents: 276
diff changeset
   448
        "v = Seq v1 (Stars vs)" "s = s1 @ s2"  "s2 \<in> (STAR r) \<rightarrow> (Stars vs)" 
42268a284ea6 updated
cu
parents: 276
diff changeset
   449
        "s1 \<in> der c r \<rightarrow> v1" "n = 0"
42268a284ea6 updated
cu
parents: 276
diff changeset
   450
         "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (STAR r))"
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   451
    (* here *)    
a3134f7de065 updated
cu
parents: 261
diff changeset
   452
    apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   453
    prefer 2
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   454
    apply(erule Posix_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   455
    apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   456
    apply(subgoal_tac "\<exists>vss. v2 = Stars vss")
a3134f7de065 updated
cu
parents: 261
diff changeset
   457
    apply(clarify)
a3134f7de065 updated
cu
parents: 261
diff changeset
   458
    apply(drule_tac x="v1" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   459
    apply(drule_tac x="vss" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   460
    apply(drule_tac x="s1" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   461
    apply(drule_tac x="s2" in meta_spec)
a3134f7de065 updated
cu
parents: 261
diff changeset
   462
     apply(simp add: der_correctness Der_def)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   463
      apply(rotate_tac 5)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   464
    apply(erule Posix_elims)
a3134f7de065 updated
cu
parents: 261
diff changeset
   465
      apply(auto)[2]
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   466
    apply(erule Posix_elims)
42268a284ea6 updated
cu
parents: 276
diff changeset
   467
      apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
   468
     apply blast
42268a284ea6 updated
cu
parents: 276
diff changeset
   469
    apply(erule Posix_elims)
42268a284ea6 updated
cu
parents: 276
diff changeset
   470
    apply(auto)
42268a284ea6 updated
cu
parents: 276
diff changeset
   471
      apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits)      
42268a284ea6 updated
cu
parents: 276
diff changeset
   472
    apply(subgoal_tac "\<exists>vss. v2 = Stars vss")
42268a284ea6 updated
cu
parents: 276
diff changeset
   473
     apply(clarify)
42268a284ea6 updated
cu
parents: 276
diff changeset
   474
    apply simp
42268a284ea6 updated
cu
parents: 276
diff changeset
   475
      apply(rotate_tac 6)
42268a284ea6 updated
cu
parents: 276
diff changeset
   476
    apply(erule Posix_elims)
42268a284ea6 updated
cu
parents: 276
diff changeset
   477
      apply(auto)[2]
42268a284ea6 updated
cu
parents: 276
diff changeset
   478
    done
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   479
    then show "(c # s) \<in> (FROMNTIMES r n) \<rightarrow> injval (FROMNTIMES r n) c v" 
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   480
    proof (cases)
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   481
      case cons
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   482
          have "s1 \<in> der c r \<rightarrow> v1" by fact
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   483
          then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   484
        moreover
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   485
          have "s2 \<in> (FROMNTIMES r (n - 1)) \<rightarrow> Stars vs" by fact
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   486
        moreover 
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   487
          have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact 
a3134f7de065 updated
cu
parents: 261
diff changeset
   488
          then have "flat (injval r c v1) = (c # s1)" by (rule Posix1)
a3134f7de065 updated
cu
parents: 261
diff changeset
   489
          then have "flat (injval r c v1) \<noteq> []" by simp
a3134f7de065 updated
cu
parents: 261
diff changeset
   490
        moreover 
a3134f7de065 updated
cu
parents: 261
diff changeset
   491
          have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (FROMNTIMES r (n - 1)))" by fact
a3134f7de065 updated
cu
parents: 261
diff changeset
   492
          then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (FROMNTIMES r (n - 1)))" 
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   493
            by (simp add: der_correctness Der_def)
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   494
        ultimately 
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   495
        have "((c # s1) @ s2) \<in> FROMNTIMES r n \<rightarrow> Stars (injval r c v1 # vs)" 
a3134f7de065 updated
cu
parents: 261
diff changeset
   496
           apply (rule_tac Posix.intros)
a3134f7de065 updated
cu
parents: 261
diff changeset
   497
               apply(simp_all)
a3134f7de065 updated
cu
parents: 261
diff changeset
   498
              apply(case_tac n)
a3134f7de065 updated
cu
parents: 261
diff changeset
   499
            apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   500
          using Posix_elims(1) FROMNTIMES.prems apply auto[1]
a3134f7de065 updated
cu
parents: 261
diff changeset
   501
          using cons(5) apply blast
a3134f7de065 updated
cu
parents: 261
diff changeset
   502
             apply(simp)
a3134f7de065 updated
cu
parents: 261
diff changeset
   503
             done
223
17c079699ea0 FROMNTIMES not yet done
Christian Urban <urbanc@in.tum.de>
parents: 222
diff changeset
   504
        then show "(c # s) \<in> FROMNTIMES r n \<rightarrow> injval (FROMNTIMES r n) c v" using cons by(simp)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   505
      next 
a3134f7de065 updated
cu
parents: 261
diff changeset
   506
       case null
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   507
          have "s1 \<in> der c r \<rightarrow> v1" by fact
42268a284ea6 updated
cu
parents: 276
diff changeset
   508
          then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp
42268a284ea6 updated
cu
parents: 276
diff changeset
   509
          moreover 
42268a284ea6 updated
cu
parents: 276
diff changeset
   510
            have "s2 \<in> STAR r \<rightarrow> Stars vs" by fact
42268a284ea6 updated
cu
parents: 276
diff changeset
   511
          moreover 
42268a284ea6 updated
cu
parents: 276
diff changeset
   512
          have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact 
42268a284ea6 updated
cu
parents: 276
diff changeset
   513
          then have "flat (injval r c v1) = (c # s1)" by (rule Posix1)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   514
          then have "flat (injval r c v1) \<noteq> []" by simp
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   515
          moreover
42268a284ea6 updated
cu
parents: 276
diff changeset
   516
             moreover 
42268a284ea6 updated
cu
parents: 276
diff changeset
   517
          have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (STAR r))" by fact
42268a284ea6 updated
cu
parents: 276
diff changeset
   518
          then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" 
42268a284ea6 updated
cu
parents: 276
diff changeset
   519
            by (simp add: der_correctness Der_def)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   520
        ultimately 
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   521
        have "((c # s1) @ s2) \<in> FROMNTIMES r 0 \<rightarrow> Stars (injval r c v1 # vs)" 
42268a284ea6 updated
cu
parents: 276
diff changeset
   522
           apply (rule_tac Posix.intros) back
42268a284ea6 updated
cu
parents: 276
diff changeset
   523
             apply(simp_all)
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   524
           done
a3134f7de065 updated
cu
parents: 261
diff changeset
   525
        then show "(c # s) \<in> FROMNTIMES r n \<rightarrow> injval (FROMNTIMES r n) c v" using null 
a3134f7de065 updated
cu
parents: 261
diff changeset
   526
          apply(simp)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   527
          done  
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   528
      qed  
a3134f7de065 updated
cu
parents: 261
diff changeset
   529
  next
a3134f7de065 updated
cu
parents: 261
diff changeset
   530
    case (NMTIMES x1 x2 m s v)
a3134f7de065 updated
cu
parents: 261
diff changeset
   531
    then show ?case sorry      
121
4c85af262ee7 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
   532
qed
106
489dfa0d7ec9 more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
   533
145
97735ef233be updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
   534
276
a3134f7de065 updated
cu
parents: 261
diff changeset
   535
section {* Lexer Correctness *}
145
97735ef233be updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
   536
151
5a1196466a9c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 150
diff changeset
   537
lemma lexer_correct_None:
145
97735ef233be updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
   538
  shows "s \<notin> L r \<longleftrightarrow> lexer r s = None"
120
d74bfa11802c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 113
diff changeset
   539
apply(induct s arbitrary: r)
143
1e7b36450d9a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 142
diff changeset
   540
apply(simp add: nullable_correctness)
120
d74bfa11802c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 113
diff changeset
   541
apply(drule_tac x="der a r" in meta_spec)
143
1e7b36450d9a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 142
diff changeset
   542
apply(auto simp add: der_correctness Der_def)
120
d74bfa11802c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 113
diff changeset
   543
done
106
489dfa0d7ec9 more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
   544
151
5a1196466a9c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 150
diff changeset
   545
lemma lexer_correct_Some:
185
841f7b9c0a6a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 172
diff changeset
   546
  shows "s \<in> L r \<longleftrightarrow> (\<exists>v. lexer r s = Some(v) \<and> s \<in> r \<rightarrow> v)"
124
5378ddbd1381 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
   547
apply(induct s arbitrary: r)
151
5a1196466a9c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 150
diff changeset
   548
apply(auto simp add: Posix_mkeps nullable_correctness)[1]
124
5378ddbd1381 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
   549
apply(drule_tac x="der a r" in meta_spec)
5378ddbd1381 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
   550
apply(simp add: der_correctness Der_def)
5378ddbd1381 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
   551
apply(rule iffI)
172
cdc0bdcfba3f updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   552
apply(auto intro: Posix_injval simp add: Posix1(1))
151
5a1196466a9c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 150
diff changeset
   553
done 
149
ec3d221bfc45 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 146
diff changeset
   554
186
0b94800eb616 added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
   555
lemma lexer_correctness:
0b94800eb616 added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
   556
  shows "(lexer r s = Some v) \<longleftrightarrow> s \<in> r \<rightarrow> v"
0b94800eb616 added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
   557
  and   "(lexer r s = None) \<longleftrightarrow> \<not>(\<exists>v. s \<in> r \<rightarrow> v)"
0b94800eb616 added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
   558
using Posix1(1) Posix_determ lexer_correct_None lexer_correct_Some apply fastforce
0b94800eb616 added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
   559
using Posix1(1) lexer_correct_None lexer_correct_Some by blast
0b94800eb616 added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
   560
95
a33d3040bf7e started a paper and moved cruft to Attic
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   561
end