thys/Re.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 30 Jan 2015 13:11:39 +0000
changeset 58 1769b702d4dc
parent 10 14d41b5b57b3
child 82 26202889f829
permissions -rw-r--r--
updated more
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
     1
(*test*)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
     2
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
     3
theory Re
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
  imports "Main" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
begin
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
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
     8
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
definition
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
  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
    11
where 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
  "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
    13
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
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
    15
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
lemma seq_empty [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
  shows "A ;; {[]} = A"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
  and   "{[]} ;; A = A"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
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
    20
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
lemma seq_null [simp]:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
  shows "A ;; {} = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
  and   "{} ;; A = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
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
    25
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
section {* Regular Expressions *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
datatype rexp =
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
  NULL
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
| EMPTY
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
| CHAR char
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
| SEQ rexp rexp
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
| ALT rexp rexp
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
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
    36
 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
  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
    39
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
  "L (NULL) = {}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
| "L (EMPTY) = {[]}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
| "L (CHAR c) = {[c]}"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
| "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
    44
| "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
    45
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    46
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    47
section {* Values *}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    48
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
datatype val = 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
  Void
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
| Char char
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
| Seq val val
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
| Right val
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
| Left val
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    56
section {* Relation between values and regular expressions *}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    57
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
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
    59
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
 "\<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
    61
| "\<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
    62
| "\<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
    63
| "\<turnstile> Void : EMPTY"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
| "\<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
    65
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
    66
thm Prf.intros(1)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
    67
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    68
section {* The string behind a value *}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    69
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
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
    71
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
  "flat(Void) = []"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
| "flat(Char c) = [c]"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
| "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
    75
| "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
    76
| "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
    77
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
lemma Prf_flat_L:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
  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
    81
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
apply(induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
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
    84
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
lemma L_flat_Prf:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
  "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
    88
apply(induct r)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
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
    90
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
    91
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
    92
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
    93
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
    94
apply (metis Prf.intros(3) flat.simps(4))
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    95
apply(erule Prf.cases)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    96
apply(auto)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    97
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
    98
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
    99
section {* Ordering of values *}
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
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
   102
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
   103
  "\<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
   104
| "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
   105
| "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
   106
| "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
   107
| "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
   108
| "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
   109
| "Void \<succ>EMPTY Void"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
| "(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
   111
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   112
(*
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
lemma
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
  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
   115
  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
   116
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
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
   124
*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   126
section {* Posix definition *}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   127
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
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
   129
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
   130
  "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
   131
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   132
(*
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
lemma POSIX_SEQ:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
  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
   135
  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
   136
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
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
   140
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
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
   142
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
   143
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
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
   145
*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
6
87618dae0e04 getting back the original version by Sulzmann
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 5
diff changeset
   147
(*
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
lemma POSIX_SEQ_I:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
  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
   150
  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
   151
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
apply(auto)
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   154
apply(rotate_tac 2)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
apply(rule ValOrd.intros)
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   159
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
done
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   162
*)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
lemma POSIX_ALT:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
  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
   166
  shows "POSIX v1 r1"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
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
   171
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
apply(erule ValOrd.cases)
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   176
apply(simp_all)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
lemma POSIX_ALT1a:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
  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
   181
  shows "POSIX v2 r2"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
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
   186
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
apply(drule mp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
apply(erule ValOrd.cases)
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   191
apply(simp_all)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   194
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
lemma POSIX_ALT1b:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
  assumes "POSIX (Right v2) (ALT r1 r2)"
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   197
  shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
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
   200
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
lemma POSIX_ALT_I1:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
  assumes "POSIX v1 r1" 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
  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
   207
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
apply(rule ValOrd.intros)
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   217
by simp
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
lemma POSIX_ALT_I2:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
  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
   221
  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
   222
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
unfolding POSIX_def
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
apply(rotate_tac 3)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
apply(rule ValOrd.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
apply metis
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
done
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   232
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   233
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   234
section {* The ordering is reflexive *}
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
lemma ValOrd_refl:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
  assumes "\<turnstile> v : r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
  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
   239
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
apply(induct)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
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
   242
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
section {* The Matcher *}
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
 nullable :: "rexp \<Rightarrow> bool"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
  "nullable (NULL) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
| "nullable (EMPTY) = True"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
| "nullable (CHAR c) = False"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
| "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
   254
| "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
   255
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
lemma nullable_correctness:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
  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
   258
apply (induct r) 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
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
   260
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
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
   263
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
  "mkeps(EMPTY) = Void"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
| "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
   266
| "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
   267
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
lemma mkeps_nullable:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
  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
   270
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
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
   272
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
   273
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
lemma mkeps_flat:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
  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
   277
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
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
   279
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   282
text {*
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   283
  The value mkeps returns is always the correct POSIX
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   284
  value.
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   285
*}
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   287
lemma mkeps_POSIX:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
  assumes "nullable r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
  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
   290
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
apply(induct r)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
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
   294
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   295
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
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
   297
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
   298
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
   299
apply(auto)[1]
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   300
apply(simp add: POSIX_def)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   301
apply(auto)[1]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   302
apply(erule Prf.cases)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   303
apply(simp_all)[5]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   304
apply(auto)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   305
apply (simp add: ValOrd.intros(2) mkeps_flat)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   306
apply(simp add: POSIX_def)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   307
apply(auto)[1]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   308
apply(erule Prf.cases)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   309
apply(simp_all)[5]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   310
apply(auto)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   311
apply (simp add: ValOrd.intros(6))
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   312
apply (simp add: ValOrd.intros(3))
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   313
apply(simp add: POSIX_def)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   314
apply(auto)[1]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   315
apply(erule Prf.cases)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   316
apply(simp_all)[5]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   317
apply(auto)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   318
apply (simp add: ValOrd.intros(6))
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   319
apply (simp add: ValOrd.intros(3))
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   320
apply(simp add: POSIX_def)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   321
apply(auto)[1]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   322
apply(erule Prf.cases)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   323
apply(simp_all)[5]
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   324
apply(auto)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   325
apply (metis Prf_flat_L mkeps_flat nullable_correctness)
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   326
by (simp add: ValOrd.intros(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
   327
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   329
section {* Derivatives *}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   330
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
fun
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
 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
   333
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
  "der c (NULL) = NULL"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
| "der c (EMPTY) = NULL"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
| "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
   337
| "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
   338
| "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
   339
     (if nullable r1
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
      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
   341
      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
   342
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
fun 
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
 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
   345
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
  "ders [] r = r"
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
| "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
   348
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   349
section {* Injection function *}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   350
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
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
   352
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
  "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
   354
| "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
   355
| "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
   356
| "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
   357
| "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
   358
| "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
   359
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   360
section {* Projection function *}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   361
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
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
   363
where
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
  "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
   365
| "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
   366
| "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
   367
| "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
   368
     (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
   369
      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
   370
                          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
   371
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   372
text {*
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   373
  Injection value is related to r
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   374
*}
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   375
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
lemma v3:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
  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
   378
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
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
   380
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
apply(simp_all)[5]
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(case_tac "c = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
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
   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(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
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
   395
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
   396
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
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
   398
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
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
   406
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
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
   408
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   409
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   410
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
apply(rule Prf.intros)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
apply(auto)[2]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   416
text {*
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   417
  The string behin the injection value is an added c
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   418
*}
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
lemma v4:
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   421
  assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
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
   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(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
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
   432
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   433
apply(auto)[1]
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(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
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
   444
apply(simp)
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(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   449
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   451
apply (metis mkeps_flat)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   454
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
done
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   457
text {*
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   458
  Injection followed by projection is the identity.
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   459
*}
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   461
lemma proj_inj_id:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
  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
   463
  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
   464
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
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
   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(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   468
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   469
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   473
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
   474
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   477
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   478
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   479
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
defer
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   484
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   485
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
   486
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   487
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   488
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   489
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   490
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   491
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   492
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   493
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
   494
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   495
apply (metis mkeps_flat)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   496
apply(auto)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   497
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   498
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   499
apply(auto)[1]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   500
apply(simp add: v4)
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
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   503
text {* 
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   504
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   505
  HERE: Crucial lemma that does not go through in the sequence case. 
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   506
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   507
*}
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   508
lemma v5:
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   509
  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
   510
  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
   511
using assms
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   512
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
   513
apply(simp)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   514
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   515
apply(simp_all)[5]
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(case_tac "c = c'")
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   521
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
   522
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   523
apply(simp_all)[5]
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   525
apply(simp_all)[5]
7
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   526
using ValOrd.simps apply blast
b409ecf47f64 cleaned up the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   527
apply(auto)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   528
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   529
apply(simp_all)[5]
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   530
(* base cases done *)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   531
(* ALT case *)
5
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   532
apply(erule Prf.cases)
fe177dfc4697 initial version of the theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   533
apply(simp_all)[5]
8
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   534
using POSIX_ALT POSIX_ALT_I1 apply blast
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   535
apply(frule POSIX_ALT1a)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   536
apply(drule POSIX_ALT1b)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   537
apply(rule POSIX_ALT_I2)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   538
apply auto[1]
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   539
apply(subst v4)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   540
apply(auto)[2]
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   541
apply(rotate_tac 1)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   542
apply(drule_tac x="v2" in meta_spec)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   543
apply(simp)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   544
apply(subst (asm) (4) POSIX_def)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   545
apply(subst (asm) v4)
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   546
apply(auto)[2]
a605dda64267 started a few arguments for the ALT case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   547
(* stuck in the ALT case *)