thys/Re.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Mon, 08 Sep 2014 16:36:13 +0100
changeset 6 87618dae0e04
parent 5 fe177dfc4697
child 7 b409ecf47f64
permissions -rw-r--r--
getting back the original version by Sulzmann
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
theory Matcher3simple
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
  imports "Main" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
begin
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
section {* Sequential Composition of Sets *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
definition
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
  Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
where 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
  "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
text {* Two Simple Properties about Sequential Composition *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
lemma seq_empty [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
  shows "A ;; {[]} = A"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
  and   "{[]} ;; A = A"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
by (simp_all add: Sequ_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
lemma seq_null [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
  shows "A ;; {} = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
  and   "{} ;; A = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
by (simp_all add: Sequ_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
section {* Regular Expressions *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
datatype rexp =
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
  NULL
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
| EMPTY
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
| CHAR char
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
| SEQ rexp rexp
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
| ALT rexp rexp
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
section {* Semantics of Regular Expressions *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
  L :: "rexp \<Rightarrow> string set"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
  "L (NULL) = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
| "L (EMPTY) = {[]}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
| "L (CHAR c) = {[c]}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
| "L (SEQ r1 r2) = (L r1) ;; (L r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
| "L (ALT r1 r2) = (L r1) \<union> (L r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
datatype val = 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
  Void
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
| Char char
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
| Seq val val
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
| Right val
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
| Left val
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
 "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
| "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
| "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
| "\<turnstile> Void : EMPTY"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
| "\<turnstile> Char c : CHAR c"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
fun flat :: "val \<Rightarrow> string"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
  "flat(Void) = []"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
| "flat(Char c) = [c]"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
| "flat(Left v) = flat(v)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
| "flat(Right v) = flat(v)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
| "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
    67
(* not actually in the paper *)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
datatype tree = 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
  Leaf string
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
| Branch tree tree
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
fun flats :: "val \<Rightarrow> tree"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
  "flats(Void) = Leaf []"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
| "flats(Char c) = Leaf [c]"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
| "flats(Left v) = flats(v)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
| "flats(Right v) = flats(v)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
| "flats(Seq v1 v2) = Branch (flats v1) (flats v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
fun flatten :: "tree \<Rightarrow> string"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
  "flatten (Leaf s) = s"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
| "flatten (Branch t1 t2) = flatten t1 @ flatten t2" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
lemma flats_flat:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
  shows "flat v1 = flatten (flats v1)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
apply(induct v1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
apply(simp_all)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
lemma Prf_flat_L:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
  assumes "\<turnstile> v : r" shows "flat v \<in> L r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
apply(induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
apply(auto simp add: Sequ_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
lemma L_flat_Prf:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
  "L(r) = {flat v | v. \<turnstile> v : r}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
apply(induct r)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
apply(auto dest: Prf_flat_L simp add: Sequ_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
apply (metis Prf.intros(4) flat.simps(1))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
apply (metis Prf.intros(5) flat.simps(2))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
apply (metis Prf.intros(1) flat.simps(5))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
apply (metis Prf.intros(2) flat.simps(3))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
apply (metis Prf.intros(3) flat.simps(4))
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   107
sorry
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   108
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   109
(*by (smt Prf.cases flat.simps(3) flat.simps(4) rexp.distinct(13) rexp.distinct(17) rexp.distinct(19) rexp.inject(3))*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
where
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   113
  "\<lbrakk>v1 = v1'; v2 \<succ>r2 v2'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   114
| "v1  \<succ>r1 v1' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
| "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
| "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
| "Void \<succ>EMPTY Void"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
| "(Char c) \<succ>(CHAR c) (Char c)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   122
(*
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
lemma
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
  assumes "r = SEQ (ALT EMPTY EMPTY) (ALT EMPTY (CHAR c))"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
  shows "(Seq (Left Void) (Right (Char c))) \<succ>r (Seq (Left Void) (Left Void))"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
done
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   134
*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
where
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   138
  "POSIX v r \<equiv> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v')"
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   140
(*
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
lemma POSIX_SEQ:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
  assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
  shows "POSIX v1 r1 \<and> POSIX v2 r2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
apply(drule_tac x="Seq v' v2" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
apply (smt Prf.intros(1) ValOrd.simps assms(3) rexp.inject(2) val.distinct(15) val.distinct(17) val.distinct(3) val.distinct(9) val.inject(2))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
apply(drule_tac x="Seq v1 v'" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
by (smt Prf.intros(1) ValOrd.simps rexp.inject(2) val.distinct(15) val.distinct(17) val.distinct(3) val.distinct(9) val.inject(2))
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   153
*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   155
(*
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
lemma POSIX_SEQ_I:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
  assumes "POSIX v1 r1" "POSIX v2 r2" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
  shows "POSIX (Seq v1 v2) (SEQ r1 r2)" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
apply(rotate_tac 4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
lemma POSIX_ALT:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
  assumes "POSIX (Left v1) (ALT r1 r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
  shows "POSIX v1 r1"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
apply(drule_tac x="Left v'" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
apply(erule ValOrd.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
lemma POSIX_ALT1a:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
  assumes "POSIX (Right v2) (ALT r1 r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
  shows "POSIX v2 r2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
apply(drule_tac x="Right v'" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
apply(erule ValOrd.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
lemma POSIX_ALT1b:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
  assumes "POSIX (Right v2) (ALT r1 r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
  shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flats v' = flats v2) \<longrightarrow> v2 \<succ>r2 v')"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
apply(drule_tac POSIX_ALT1a)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
lemma POSIX_ALT_I1:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
  assumes "POSIX v1 r1" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
  shows "POSIX (Left v1) (ALT r1 r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
by (metis flats_flat order_refl)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
lemma POSIX_ALT_I2:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
  assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
  shows "POSIX (Right v2) (ALT r1 r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
done
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   240
*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
lemma ValOrd_refl:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
  assumes "\<turnstile> v : r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
  shows "v \<succ>r v"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
apply(induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
apply(auto intro: ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   250
(*
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
lemma ValOrd_length:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
  assumes "v1 \<succ>r v2" shows "length (flat v1) \<ge> length (flat v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
apply(induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
done
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   257
*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
section {* The Matcher *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
 nullable :: "rexp \<Rightarrow> bool"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
  "nullable (NULL) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
| "nullable (EMPTY) = True"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
| "nullable (CHAR c) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
lemma nullable_correctness:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
apply (induct r) 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
apply(auto simp add: Sequ_def) 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
fun mkeps :: "rexp \<Rightarrow> val"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
  "mkeps(EMPTY) = Void"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
lemma mkeps_nullable:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
  assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
apply(induct rule: nullable.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
apply(auto intro: Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   287
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
lemma mkeps_flat:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
  assumes "nullable(r)" shows "flat (mkeps r) = []"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
apply(induct rule: nullable.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   295
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   297
lemma mkeps_flats:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   298
  assumes "nullable(r)" shows "flatten (flats (mkeps r)) = []"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   299
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
apply(induct rule: nullable.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   301
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
lemma mkeps_POSIX:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
  assumes "nullable r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
  shows "POSIX (mkeps r) r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
apply(induct r)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   311
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   313
apply(simp_all)[5]
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   314
apply (metis ValOrd.intros)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
apply(auto)[1]
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   317
sorry
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   318
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
  "der c (NULL) = NULL"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
| "der c (EMPTY) = NULL"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
| "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
| "der c (SEQ r1 r2) = 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
     (if nullable r1
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
      then ALT (SEQ (der c r1) r2) (der c r2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
      else SEQ (der c r1) r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
fun 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
 ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
  "ders [] r = r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
| "ders (c # s) r = ders s (der c r)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
  "injval (CHAR d) c Void = Char d"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
| "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
| "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
| "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
| "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
| "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
  "projval (CHAR d) c _ = Void"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
| "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
| "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
| "projval (SEQ r1 r2) c (Seq v1 v2) = 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
     (if flat v1 = [] then Right(projval r2 c v2) 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
      else if nullable r1 then Left (Seq (projval r1 c v1) v2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
                          else Seq (projval r1 c v1) v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
lemma v3:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
  assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
apply(induct arbitrary: v rule: der.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
apply(case_tac "c = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
apply (metis Prf.intros(5))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   373
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
apply (metis Prf.intros(2))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
apply (metis Prf.intros(3))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
apply(case_tac "nullable r1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   385
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
apply (metis Prf.intros(1))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
apply (metis Prf.intros(1) mkeps_nullable)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
apply(auto)[2]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   396
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
fun head where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
  "head [] = None"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
| "head (x#xs) = Some x"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
lemma head1:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
  assumes "head (xs @ ys) = Some x"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
  shows "(head xs = Some x) \<or> (xs = [] \<and> head ys = Some x)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
apply(induct xs)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   408
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   409
lemma head2:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   410
  assumes "head (xs @ ys) = None"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
  shows "(head xs = None) \<and> (head ys = None)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
apply(induct xs)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   416
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   417
lemma v4:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   418
  assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c#(flat v)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
apply(induct arbitrary: v rule: der.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   421
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   426
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   427
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
apply(case_tac "c = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   432
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   433
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   437
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   438
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   440
apply(case_tac "nullable r1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   444
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   445
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   447
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   448
apply (metis mkeps_flat)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   449
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   451
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   454
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
lemma proj_inj_id:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
  assumes "\<turnstile> v : der c r" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   457
  shows "projval r c (injval r c v) = v"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   458
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   459
apply(induct r arbitrary: c v rule: rexp.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   461
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   463
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   464
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   467
apply(case_tac "c = char")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   468
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   469
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   473
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   474
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   477
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   478
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   479
apply(case_tac "nullable rexp1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   484
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   485
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   486
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   487
apply (metis list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   488
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   489
apply (metis mkeps_flat)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   490
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   491
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   492
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   493
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   494
apply(simp add: v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   495
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   496
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   497
lemma Le_2a: "(head (flat v) = None) = (flat v = [])"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   498
apply(induct v)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   499
apply(simp_all)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   500
apply (metis Nil_is_append_conv head.elims option.distinct(1))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   501
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   502
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   503
(* HERE *)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   504
lemma v5:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   505
  assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   506
  shows "POSIX (injval r c v) r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   507
using assms
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   508
apply(induct r)
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   509
apply(simp (no_asm) only: POSIX_def)
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   510
apply(rule allI)
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   511
apply(rule impI)
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   512
apply(simp)
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   513
apply(erule Prf.cases)
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   514
apply(simp_all)[5]
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   515
apply(induct arbitrary: v rule: der.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   516
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   517
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   518
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   519
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   520
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   521
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   522
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   523
apply(case_tac "c = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
apply(auto simp add: POSIX_def)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   525
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   526
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   527
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   528
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   529
apply (metis ValOrd.intros(7))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   530
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   531
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   532
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   533
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   534
apply(case_tac "nullable r1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   535
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   536
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   537
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   538
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   539
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   540
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   541
apply(drule POSIX_SEQ)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   542
apply(assumption)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   543
apply(assumption)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   544
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   545
apply(rule POSIX_SEQ_I)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   546
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   547
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   548
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   549
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   550
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   551
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   552
apply(drule POSIX_ALT)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   553
apply(rule POSIX_ALT_I1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   554
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   555
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   556
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   557
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   558
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   559
apply(rotate_tac 5)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   560
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   561
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   562
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   563
apply(drule POSIX_ALT)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   564
apply(drule POSIX_SEQ)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   565
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   566
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   567
apply(rule POSIX_SEQ_I)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   568
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   569
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   570
apply(drule POSIX_ALT1a)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   571
apply(rule POSIX_SEQ_I)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   572
apply (metis mkeps_POSIX)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   573
apply(metis)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   574
apply(frule POSIX_ALT1a)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   575
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   576
apply(drule_tac x="v2" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   577
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   578
apply(rule POSIX_ALT_I2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   579
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   580
apply(frule POSIX_ALT1a)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   581
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   582
apply(subst v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   583
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   584
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   585
apply(case_tac "\<exists>r. v2 \<succ>r v'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   586
apply (metis ValOrd_length le_neq_implies_less less_Suc_eq)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   587
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   588
apply(frule_tac x="der c r2" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   589
apply(subgoal_tac "\<not>(Right v2 \<succ>(ALT (der c r1) (der c r2)) Right v')")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   590
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   591
apply(rule notI)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   592
apply(erule ValOrd.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   593
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   594
apply(subst (asm) (1) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   595
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   596
apply(subgoal_tac "(\<not> \<turnstile> Right v' : ALT (der c r1) (der c r2)) \<or> (flats v2 \<noteq> flats v')")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   597
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   598
apply (metis flats.simps(4))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   599
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   600
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   601
apply(subst v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   602
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   603
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   604
apply(case_tac "v2 \<succ>r1 v'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   605
apply (metis ValOrd_length le_neq_implies_less less_Suc_eq)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   606
apply(
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   607
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   608
apply(drule_tac x="projval (ALT r1 r2) c v'" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   609
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   610
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   611
apply(frule POSIX_ALT1b)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   612
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   613
apply(subst v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   614
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   615
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   616
apply(simp add: flats_flat)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   617
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   618
apply(drule_tac x="v2" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   619
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   620
apply(rule POSIX_ALT_I2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   621
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   622
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   623
apply(subst v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   624
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   625
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   626
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   627
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   628
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   629
apply(rotate_tac 5)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   630
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   631
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   632
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   633
apply(drule POSIX_ALT)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   634
apply(drule POSIX_SEQ)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   635
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   636
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   637
apply(rule POSIX_SEQ_I)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   638
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   639
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   640
apply(drule POSIX_ALT1a)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   641
apply(rule POSIX_SEQ_I)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   642
apply (metis mkeps_POSIX)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   643
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   644
apply(drule_tac x="projval r1 c v'" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   645
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   646
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   647
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   648
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   649
apply(subst (asm) (1) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   650
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   651
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   652
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   653
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   654
lemma inj_proj_id:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   655
  assumes "\<turnstile> v : r" "POSIX v r" "head (flat v) = Some c"  
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   656
  shows "injval r c (projval r c v) = v"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   657
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   658
apply(induct v r arbitrary: rule: Prf.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   659
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   660
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   661
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   662
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   663
apply(frule POSIX_head)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   664
apply(assumption)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   665
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   666
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   667
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   668
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   669
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   673
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   674
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   675
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   676
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   677
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   678
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   679
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   680
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   681
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   682
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   683
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   684
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   685
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   686
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   687
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   688
apply(case_tac "c = char")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   689
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   690
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   691
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   692
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   693
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   694
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   695
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   696
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   697
lemma POSIX_head:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   698
  assumes "POSIX (Stars (v#vs)) (STAR r)" "head (flat v @ flat (Stars vs)) = Some c"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   699
  shows "head (flat v) = Some c"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   700
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   701
apply(rule_tac ccontr)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   702
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   703
apply(drule Le_1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   704
apply(assumption)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   705
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   706
apply(auto simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   707
apply(drule_tac x="Stars (v'#vs')" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   708
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   709
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   710
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   711
apply(erule ValOrd.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   712
apply(simp_all)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   713
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   714
using Le_2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   715
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   716
by (metis Le_2 Le_2a head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   717
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   718
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   719
lemma inj_proj_id:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   720
  assumes "\<turnstile> v : r" "POSIX v r" "head (flat v) = Some c"  
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   721
  shows "injval r c (projval r c v) = v"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   722
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   723
apply(induct r arbitrary: rule: Prf.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   724
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   725
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   726
apply(frule POSIX_head)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   727
apply(assumption)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   728
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   729
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   730
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   731
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   732
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   733
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   734
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   735
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   736
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   737
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   738
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   739
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   740
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   741
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   742
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   743
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   744
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   745
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   746
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   747
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   748
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   749
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   750
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   751
apply(case_tac "c = char")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   752
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   753
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   754
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   755
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   756
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   757
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   758
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   759
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   760
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   761
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   762
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   763
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   764
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   765
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   766
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   767
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   768
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   769
apply(case_tac "nullable rexp1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   770
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   771
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   772
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   773
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   774
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   775
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   776
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   777
apply(simp add: v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   778
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   779
apply(simp add: v2a)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   780
apply(simp only: der.simps)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   781
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   782
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   783
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   784
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   785
apply(simp add: v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   786
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   787
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   788
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   789
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   790
lemma STAR_obtain:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   791
  assumes "\<turnstile> v : STAR r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   792
  obtains vs where "\<turnstile> Stars vs : STAR r" and "v = Stars vs"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   793
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   794
by (smt Prf.cases rexp.distinct(17) rexp.distinct(23) rexp.distinct(27) rexp.distinct(29))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   795
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   796
fun first :: "val \<Rightarrow> char option"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   797
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   798
  "first Void = None"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   799
| "first (Char c) = Some c"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   800
| "first (Seq v1 v2) = (if (\<exists>c. first v1 = Some c) then first v1 else first v2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   801
| "first (Right v) = first v"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   802
| "first (Left v) = first v"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   803
| "first (Stars []) = None"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   804
| "first (Stars (v#vs)) =  (if (\<exists>c. first v = Some c) then first v else first (Stars vs))"   
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   805
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
lemma flat: 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   807
  shows "flat v = [] \<longleftrightarrow> first v = None"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   808
  and "flat (Stars vs) = [] \<longleftrightarrow> first (Stars vs) = None"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   809
apply(induct v and vs)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   810
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   811
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   812
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   813
lemma first: 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   814
  shows "first v = Some c \<Longrightarrow> head (flat v @ ys) = Some c"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   815
apply(induct arbitrary: ys rule: first.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   816
apply(auto split: if_splits simp add: flat[symmetric])
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   817
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   818
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   819
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   820
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   821
lemma v5:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   822
  assumes "POSIX v r" "head (flat v) = Some c" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   823
  shows "\<turnstile> projval r c v : der c r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   824
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   825
apply(induct r arbitrary: c v rule: rexp.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   826
prefer 6
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   827
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   828
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   829
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   830
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   831
apply(drule_tac x="c" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   832
apply(drule_tac x="va" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   833
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   834
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   835
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   836
apply(drule_tac x="Stars (v'#vs)" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   837
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   838
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   839
apply (metis Prf.intros(2))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   840
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   841
apply(erule disjE)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   842
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   843
apply(erule ValOrd.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   844
apply(simp_all)[10]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   845
apply(drule_tac meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   846
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   847
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   848
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   849
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   850
apply(drule_tac x="Stars (va#vs)" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   851
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   852
prefer 6
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   853
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   854
apply(auto split: if_splits)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   855
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   856
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   857
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   858
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   859
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   860
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   861
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   862
apply(auto split: if_splits)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   863
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   864
apply(auto split: if_splits)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   865
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   866
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   867
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   868
apply(case_tac "char = c")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   869
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   870
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   871
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   872
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   873
apply(simp_all)[
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   874
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   875
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   876
lemma uu:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   877
  assumes ih: "\<And>v c. \<lbrakk>head (flat v) = Some c\<rbrakk> \<Longrightarrow> \<turnstile> projval rexp c v : der c rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   878
  assumes "\<turnstile> Stars vs : STAR rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   879
  and "first (Stars (v#vs)) = Some c"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   880
  shows "\<turnstile> projval rexp c v : der c rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   881
using assms(2,3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   882
apply(induct vs)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   883
apply(simp_all)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   884
apply(auto split: if_splits)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   885
apply (metis first head1 ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   886
apply (metis first head1 ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   887
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   888
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   889
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   890
lemma v5:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   891
  assumes "\<turnstile> v : r" "head (flat v) = Some c" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   892
  shows "\<turnstile> projval r c v : der c r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   893
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   894
apply(induct r arbitrary: v c rule: rexp.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   895
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   896
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   897
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   898
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   899
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   900
apply(case_tac "char = c")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   901
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   902
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   903
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   904
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   905
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   906
prefer 3
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   907
proof -
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   908
  fix rexp v c vs
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   909
  assume ih: "\<And>v c. \<lbrakk>\<turnstile> v : rexp; head (flat v) = Some c\<rbrakk> \<Longrightarrow> \<turnstile> projval rexp c v : der c rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   910
  assume a: "head (flat (Stars vs)) = Some c"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   911
  assume b: "\<turnstile> Stars vs : STAR rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   912
  have c: "first (Stars vs) = Some c \<Longrightarrow> \<turnstile> (hd vs) : rexp \<Longrightarrow> \<turnstile> projval rexp c (hd vs) : der c rexp"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   913
    using b
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   914
    apply(induct arbitrary: rule: Prf.induct) 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   915
    apply(simp_all)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   916
    apply(case_tac "\<exists>c. first v = Some c")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   917
    apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   918
    apply(rule ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   919
    apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   920
    apply (metis append_Nil2 first)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   921
    apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   922
    apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   923
    apply(auto split: if_splits)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   924
apply (metis append_Nil2 first ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   925
    apply(drule_tac x="v" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   926
    apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   927
    apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   928
    apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   929
    apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   930
    apply(drule_tac x="v" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   931
    apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   932
    using a
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   933
    apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   934
    apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   935
    apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   936
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   937
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   938
apply (metis (full_types) Prf.cases Prf.simps a b flat.simps(6) head.simps(1) list.distinct(1) list.inject option.distinct(1) rexp.distinct(17) rexp.distinct(23) rexp.distinct(27) rexp.distinct(29) val.distinct(17) val.distinct(27) val.distinct(29) val.inject(5))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   939
    
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   940
    apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   941
    apply(case_tac "\<exists>c. first a = Some c")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   942
    apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   943
    apply(rule ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   944
    apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   945
apply (metis append_Nil2 first ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   946
    apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   947
    apply(case_tac "\<exists>c. first a = Some c")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   948
    apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   949
    apply(rule ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   950
    apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   951
    apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   952
    
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   953
    apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   954
    apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   955
apply (metis append_Nil2 first)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   956
    apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   957
    apply(simp add: )
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   958
    apply(simp add: flat)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   959
    apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   960
    apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   961
    apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   962
      *)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   963
  show "\<turnstile> projval (STAR rexp) c (Stars vs) : SEQ (der c rexp) (STAR rexp)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   964
    using b a
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   965
    apply -
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   966
    apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   967
    apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   968
    apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   969
    apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   970
    apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   971
    apply(rule ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   972
    apply(simp_all)[3]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   973
    using k
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   974
    apply -
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   975
    apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   976
    apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   977
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   978
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   979
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   980
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   981
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   982
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   983
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   984
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   985
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   986
apply(drule_tac meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   987
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   988
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   989
apply(rule ih)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   990
apply(case_tac vs)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   991
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   992
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   993
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   994
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   995
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   996
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   997
apply (metis Prf.intros(1) Prf.intros(3))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   998
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   999
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1000
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1001
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1002
apply (metis Prf.intros(3))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1003
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1004
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1005
apply (metis Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1006
apply(case_tac "nullable r1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1007
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1008
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1009
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1010
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1011
apply (metis Prf.intros(5))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1012
apply (metis Prf.intros(3) Prf.intros(4) head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1013
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1014
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1015
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1016
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1017
apply (metis nullable_correctness v1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1018
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1019
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1020
apply (metis Prf.intros(3))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1021
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1022
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1023
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1024
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1025
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1026
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1027
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1028
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1029
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1030
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1031
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1032
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1033
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1034
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1035
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1036
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1037
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1038
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1039
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1040
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1041
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1042
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1043
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1044
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1045
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1046
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1047
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1048
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1049
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1050
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1051
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1052
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1053
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1054
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1055
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1056
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1057
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1058
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1059
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1060
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1061
apply(drule_tac x="v" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1062
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1063
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1064
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1065
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1066
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1067
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1068
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1069
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1070
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1071
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1072
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1073
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1074
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1075
apply (metis Prf.intros(3) head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1076
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1077
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1078
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1079
apply(drule_tac head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1080
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1081
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1082
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1083
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1084
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1085
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1086
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1087
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1088
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1089
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1090
apply (metis nullable_correctness v1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1091
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1092
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1093
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1094
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1095
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1096
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1097
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1098
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1099
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1100
apply(drule_tac head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1101
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1102
apply (metis Prf.intros(3))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1103
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1104
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1105
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1106
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1107
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1108
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1109
apply(drule_tac x="c" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1110
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1111
apply(rotate_tac 6)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1112
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1113
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1114
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1115
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1116
apply(auto)[2]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1117
apply(drule v1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1118
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1119
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1120
apply(simp add: nullable_correctness[symmetric])
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1121
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1122
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1123
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1124
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1125
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1126
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1127
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1128
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1129
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1130
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1131
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1132
lemma inj_proj_id:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1133
  assumes "\<turnstile> v : r" "head (flat v) = Some c" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1134
  shows "injval r c (projval r c v) = v"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1135
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1136
apply(induct arbitrary: c rule: Prf.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1137
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1138
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1139
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1140
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1141
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1142
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1143
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1144
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1145
apply(drule head1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1146
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1147
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1148
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1149
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1150
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1151
apply(case_tac "ca = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1152
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1153
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1154
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1155
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1156
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1157
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1158
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1159
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1160
apply(case_tac "nullable r1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1161
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1162
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1163
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1164
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1165
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1166
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1167
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1168
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1169
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1170
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1171
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1172
apply(case_tac "nullable rexp1")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1173
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1174
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1175
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1176
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1177
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1178
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1179
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1180
apply(simp add: v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1181
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1182
apply(simp add: v2a)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1183
apply(simp only: der.simps)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1184
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1185
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1186
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1187
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1188
apply(simp add: v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1189
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1190
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1191
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1192
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1193
lemma POSIX_I:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1194
  assumes "\<turnstile> v : r" "\<And>v'.  \<turnstile> v' : r \<Longrightarrow> flat v = flat v' \<Longrightarrow> v \<succeq>r v'"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1195
  shows "POSIX v r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1196
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1197
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1198
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1199
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1200
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1201
lemma DISJ:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1202
  "(\<not> A \<Longrightarrow> B) \<Longrightarrow> A \<or> B"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1203
by metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1204
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1205
lemma DISJ2:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1206
  "\<not>(A \<and> B) \<longleftrightarrow> \<not>A \<or> \<not>B"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1207
by metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1208
 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1209
lemma APP: "xs @ [] = xs" by simp
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1210
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1211
lemma v5:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1212
  assumes "POSIX v (der c r)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1213
  shows "POSIX (injval r c v) r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1214
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1215
apply(induct arbitrary: v rule: der.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1216
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1217
apply(auto simp add: POSIX_def)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1218
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1219
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1220
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1221
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1222
apply(auto simp add: POSIX_def)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1223
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1224
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1225
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1226
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1227
apply(case_tac "c = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1228
apply(auto simp add: POSIX_def)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1229
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1230
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1231
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1232
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1233
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1234
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1235
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1236
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1237
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1238
apply(auto simp add: POSIX_def)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1239
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1240
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1241
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1242
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1243
prefer 3
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1244
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1245
apply(frule POSIX_E1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1246
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1247
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1248
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1249
apply(drule_tac x="v1" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1250
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1251
apply(rule POSIX_I)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1252
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1253
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1254
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1255
apply(drule_tac x="Seq v' v2" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1256
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1257
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1258
apply (metis Prf.intros(3))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1259
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1260
apply(erule disjE)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1261
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1262
apply(erule ValOrd.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1263
apply(simp_all)[10]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1264
apply(subgoal_tac "\<exists>vs2. v2 = Stars vs2")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1265
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1266
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1267
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1268
apply(auto)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1269
apply(erule exE)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1270
apply(clarify)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1271
apply(simp only: injval.simps)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1272
apply(case_tac "POSIX (Stars (injval r c v1 # vs2)) (STAR r)")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1273
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1274
apply(subgoal_tac "\<exists>v'' vs''. Stars (v''#vs'')  \<succ>(STAR r) Stars (injval r c v1 # vs2)")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1275
apply(erule exE)+
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1276
apply(erule ValOrd.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1277
apply(simp_all)[8]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1278
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1279
apply(subst (asm) (2) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1280
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1281
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1282
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1283
apply(subst (asm) (3) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1284
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1285
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1286
apply (metis Prf.intros(2) v3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1287
apply(erule exE)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1288
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1289
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1290
apply
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1291
apply(subgoal_tac "\<turnstile> Stars (injval r c v1 # vs2) : STAR r")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1292
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1293
apply(case_tac "flat () = []")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1294
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1295
apply(rule POSIX_I)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1296
apply (metis POSIX_E1 der.simps(6) v3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1297
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1298
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1299
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1300
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1301
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1302
prefer 4
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1303
apply(clarify)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1304
apply(simp only: v4 flat.simps injval.simps)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1305
prefer 4
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1306
apply(clarify)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1307
apply(simp only: v4 APP flat.simps injval.simps)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1308
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1309
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1310
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1311
apply(clarify)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1312
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1313
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1314
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1315
apply(subst (asm) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1316
apply(erule conjE)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1317
apply(drule_tac x="va" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1318
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1319
apply(simp add: v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1320
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1321
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1322
apply(rule disjI2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1323
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1324
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1325
apply(subst (asm) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1326
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1327
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1328
apply(drule_tac x="v" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1329
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1330
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1331
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1332
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1333
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1334
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1335
apply(case_tac "injval rb c v1 = v \<and> [] = vs")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1336
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1337
apply(simp only: DISJ2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1338
apply(rule disjI2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1339
apply(erule disjE)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1340
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1341
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1342
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1343
apply (metis list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1344
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1345
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1346
apply(rule DISJ)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1347
apply(simp only: DISJ2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1348
apply(erule disjE)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1349
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1350
apply(subst (asm) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1351
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1352
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1353
apply (metis list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1354
apply(subst (asm) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1355
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1356
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1357
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1358
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1359
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1360
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1361
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1362
apply (metis list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1363
apply (metis list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1364
apply(clarify)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1365
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1366
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1367
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1368
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1369
apply(clarify)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1370
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1371
apply(drule_tac x="Stars (v#vs)" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1372
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1373
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1374
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1375
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1376
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1377
apply (metis POSIX_E1 der.simps(6) list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1378
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1379
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1380
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1381
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1382
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1383
apply(case_tac v2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1384
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1385
apply(simp (no_asm) POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1386
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1387
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1388
 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1389
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1390
apply (metis POSIX_E der.simps(6) v3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1391
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1392
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1393
apply(drule_tac x="v1" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1394
apply(drule meta_mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1395
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1396
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1397
apply(subgoal_tac "POSIX v1 (der c r)")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1398
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1399
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1400
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1401
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1402
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1403
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1404
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1405
lemma v5:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1406
  assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1407
  shows "POSIX (injval r c v) r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1408
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1409
apply(induct arbitrary: v rule: der.induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1410
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1411
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1412
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1413
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1414
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1415
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1416
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1417
apply(case_tac "c = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1418
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1419
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1420
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1421
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1422
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1423
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1424
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1425
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1426
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1427
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1428
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1429
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1430
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1431
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1432
prefer 3
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1433
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1434
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1435
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1436
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1437
apply(case_tac "flat (Seq v1 v2) \<noteq> []")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1438
apply(subgoal_tac "POSIX v1 (der c r)")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1439
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1440
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1441
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1442
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1443
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1444
apply(drule_tac x="v1" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1445
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1446
apply(simp (no_asm) add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1447
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1448
apply(rule v3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1449
apply (metis Prf.intros(3) der.simps(6))
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1450
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1451
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1452
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1453
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1454
apply(rule disjI2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1455
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1456
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1457
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1458
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1459
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1460
apply (metis list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1461
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1462
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1463
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1464
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1465
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1466
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1467
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1468
apply(drule_tac x="v" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1469
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1470
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1471
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1472
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1473
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1474
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1475
apply(
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1476
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1477
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1478
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1479
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1480
apply (metis list.distinct(1) v4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1481
apply(rotate_tac 8)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1482
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1483
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1484
apply (metis POSIX_def ValOrd.intros(9) ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1485
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1486
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1487
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1488
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1489
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1490
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1491
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1492
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1493
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1494
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1495
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1496
apply(erule conjE)+
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1497
apply(drule_tac x="Seq v1 v2" in spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1498
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1499
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1500
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1501
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1502
apply(rotate_tac 2)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1503
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1504
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1505
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1506
apply(simp add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1507
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1508
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1509
apply(rule v3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1510
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1511
apply(rotate_tac 5)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1512
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1513
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1514
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1515
apply(rotate_tac 4)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1516
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1517
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1518
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1519
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1520
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1521
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1522
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1523
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1524
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1525
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1526
apply(subgoal_tac "POSIX v2 (der c r2)")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1527
apply(rotate_tac 1)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1528
apply(drule_tac x="v2" in meta_spec)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1529
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1530
apply(subgoal_tac "\<turnstile> Right (injval r2 c v2) : (ALT r1 r2)") 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1531
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1532
apply (metis Prf.intros(5) v3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1533
apply(simp (no_asm) add: POSIX_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1534
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1535
apply(rotate_tac 6)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1536
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1537
apply(simp_all)[7]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1538
prefer 2
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1539
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1540
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1541
apply (metis POSIX_def ValOrd.intros(5) ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1542
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1543
apply(simp add: ValOrd_eq_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1544
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1545
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1546
section {* Correctness Proof of the Matcher *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1547
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1548
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1549
section {* Left-Quotient of a Set *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1550
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1551
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1552
 zeroable :: "rexp \<Rightarrow> bool"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1553
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1554
  "zeroable (NULL) = True"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1555
| "zeroable (EMPTY) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1556
| "zeroable (CHAR c) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1557
| "zeroable (ALT r1 r2) = (zeroable r1 \<and> zeroable r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1558
| "zeroable (SEQ r1 r2) = (zeroable r1 \<or> zeroable r2)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1559
| "zeroable (STAR r) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1560
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1561
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1562
lemma zeroable_correctness:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1563
  shows "zeroable r  \<longleftrightarrow>  (L r = {})"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1564
apply(induct r)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1565
apply(auto simp add: Seq_def)[6]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1566
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1567
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1568
section {* Left-Quotient of a Set *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1569
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1570
definition
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1571
  Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1572
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1573
  "Der c A \<equiv> {s. [c] @ s \<in> A}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1574
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1575
lemma Der_null [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1576
  shows "Der c {} = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1577
unfolding Der_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1578
by auto
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1579
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1580
lemma Der_empty [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1581
  shows "Der c {[]} = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1582
unfolding Der_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1583
by auto
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1584
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1585
lemma Der_char [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1586
  shows "Der c {[d]} = (if c = d then {[]} else {})"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1587
unfolding Der_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1588
by auto
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1589
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1590
lemma Der_union [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1591
  shows "Der c (A \<union> B) = Der c A \<union> Der c B"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1592
unfolding Der_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1593
by auto
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1594
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1595
lemma Der_seq [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1596
  shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1597
unfolding Der_def Seq_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1598
by (auto simp add: Cons_eq_append_conv)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1599
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1600
lemma Der_star [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1601
  shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1602
proof -    
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1603
  have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1604
    by (simp only: star_cases[symmetric])
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1605
  also have "... = Der c (A ;; A\<star>)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1606
    by (simp only: Der_union Der_empty) (simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1607
  also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1608
    by simp
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1609
  also have "... =  (Der c A) ;; A\<star>"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1610
    unfolding Seq_def Der_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1611
    by (auto dest: star_decomp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1612
  finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1613
qed
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1614
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1615
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1616
lemma der_correctness:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1617
  shows "L (der c r) = Der c (L r)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1618
by (induct r) 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1619
   (simp_all add: nullable_correctness)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1620
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1621
lemma matcher_correctness:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1622
  shows "matcher r s \<longleftrightarrow> s \<in> L r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1623
by (induct s arbitrary: r)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1624
   (simp_all add: nullable_correctness der_correctness Der_def)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1625
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1626
section {* Examples *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1627
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1628
definition 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1629
  "CHRA \<equiv> CHAR (CHR ''a'')"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1630
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1631
definition 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1632
  "ALT1 \<equiv> ALT CHRA EMPTY"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1633
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1634
definition 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1635
  "SEQ3 \<equiv> SEQ (SEQ ALT1 ALT1) ALT1"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1636
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1637
value "matcher SEQ3 ''aaa''"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1638
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1639
value "matcher NULL []"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1640
value "matcher (CHAR (CHR ''a'')) [CHR ''a'']"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1641
value "matcher (CHAR a) [a,a]"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1642
value "matcher (STAR (CHAR a)) []"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1643
value "matcher (STAR (CHAR a))  [a,a]"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1644
value "matcher (SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ''abbbbc''"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1645
value "matcher (SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ''abbcbbc''"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1646
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1647
section {* Incorrect Matcher - fun-definition rejected *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1648
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1649
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1650
  match :: "rexp list \<Rightarrow> string \<Rightarrow> bool"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1651
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1652
  "match [] [] = True"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1653
| "match [] (c # s) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1654
| "match (NULL # rs) s = False"  
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1655
| "match (EMPTY # rs) s = match rs s"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1656
| "match (CHAR c # rs) [] = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1657
| "match (CHAR c # rs) (d # s) = (if c = d then match rs s else False)"         
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1658
| "match (ALT r1 r2 # rs) s = (match (r1 # rs) s \<or> match (r2 # rs) s)" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1659
| "match (SEQ r1 r2 # rs) s = match (r1 # r2 # rs) s"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1660
| "match (STAR r # rs) s = (match rs s \<or> match (r # (STAR r) # rs) s)"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1661
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1662
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1663
end