thys/Re1.thy~
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Mon, 06 Oct 2014 15:24:41 +0100
changeset 20 c11651bbebf5
parent 12 232ea10bed6f
permissions -rw-r--r--
some small changes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     1
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     2
theory Re1
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     3
  imports "Main" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     4
begin
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     5
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     6
section {* Sequential Composition of Sets *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     7
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     8
definition
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     9
  Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    10
where 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    11
  "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    13
text {* Two Simple Properties about Sequential Composition *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    14
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    15
lemma seq_empty [simp]:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    16
  shows "A ;; {[]} = A"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    17
  and   "{[]} ;; A = A"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    18
by (simp_all add: Sequ_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    19
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    20
lemma seq_null [simp]:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    21
  shows "A ;; {} = {}"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    22
  and   "{} ;; A = {}"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    23
by (simp_all add: Sequ_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    24
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    25
section {* Regular Expressions *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    26
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    27
datatype rexp =
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    28
  NULL
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    29
| EMPTY
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    30
| CHAR char
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    31
| SEQ rexp rexp
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    32
| ALT rexp rexp
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    33
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    34
section {* Semantics of Regular Expressions *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    35
 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    36
fun
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    37
  L :: "rexp \<Rightarrow> string set"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    38
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    39
  "L (NULL) = {}"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    40
| "L (EMPTY) = {[]}"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    41
| "L (CHAR c) = {[c]}"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    42
| "L (SEQ r1 r2) = (L r1) ;; (L r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    43
| "L (ALT r1 r2) = (L r1) \<union> (L r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    44
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    45
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    46
section {* Values *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    47
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    48
datatype val = 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    49
  Void
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    50
| Char char
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    51
| Seq val val
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    52
| Right val
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    53
| Left val
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    54
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    55
section {* Relation between values and regular expressions *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    56
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    57
inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    58
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    59
 "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    60
| "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    61
| "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    62
| "\<turnstile> Void : EMPTY"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    63
| "\<turnstile> Char c : CHAR c"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    64
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    65
section {* The string behind a value *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    66
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    67
fun flat :: "val \<Rightarrow> string"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    68
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    69
  "flat(Void) = []"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    70
| "flat(Char c) = [c]"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    71
| "flat(Left v) = flat(v)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    72
| "flat(Right v) = flat(v)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    73
| "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    74
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    75
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    76
lemma Prf_flat_L:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    77
  assumes "\<turnstile> v : r" shows "flat v \<in> L r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    78
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    79
apply(induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    80
apply(auto simp add: Sequ_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    81
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    82
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    83
lemma L_flat_Prf:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    84
  "L(r) = {flat v | v. \<turnstile> v : r}"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    85
apply(induct r)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    86
apply(auto dest: Prf_flat_L simp add: Sequ_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    87
apply (metis Prf.intros(4) flat.simps(1))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    88
apply (metis Prf.intros(5) flat.simps(2))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    89
apply (metis Prf.intros(1) flat.simps(5))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    90
apply (metis Prf.intros(2) flat.simps(3))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    91
apply (metis Prf.intros(3) flat.simps(4))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    92
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    93
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    94
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    95
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    96
section {* Ordering of values *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    97
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    98
inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    99
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   100
  "\<lbrakk>v1 = v1'; v2 \<succ>r2 v2'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   101
| "v1  \<succ>r1 v1' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   102
| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   103
| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   104
| "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   105
| "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   106
| "Void \<succ>EMPTY Void"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   107
| "(Char c) \<succ>(CHAR c) (Char c)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   108
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   109
(*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   110
lemma
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   111
  assumes "r = SEQ (ALT EMPTY EMPTY) (ALT EMPTY (CHAR c))"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   112
  shows "(Seq (Left Void) (Right (Char c))) \<succ>r (Seq (Left Void) (Left Void))"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   113
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   114
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   115
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   116
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   117
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   118
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   119
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   120
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   121
*)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   122
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   123
section {* Posix definition *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   124
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   125
definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   126
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   127
  "POSIX v r \<equiv> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v')"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   128
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   129
(*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   130
an alternative definition: might cause problems
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   131
with theorem mkeps_POSIX
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   132
*)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   133
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   134
definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   135
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   136
  "POSIX2 v r \<equiv> \<turnstile> v : r \<and> (\<forall>v'. \<turnstile> v' : r \<longrightarrow> v \<succ>r v')"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   137
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   138
definition POSIX3 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   139
where
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   140
  "POSIX3 v r \<equiv> \<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v \<ge> flat v') \<longrightarrow> v \<succ>r v')"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   141
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   142
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   143
(*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   144
lemma POSIX_SEQ:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   145
  assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   146
  shows "POSIX v1 r1 \<and> POSIX v2 r2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   147
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   148
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   149
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   150
apply(drule_tac x="Seq v' v2" in spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   151
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   152
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))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   153
apply(drule_tac x="Seq v1 v'" in spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   154
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   155
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))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   156
*)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   157
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   158
(*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   159
lemma POSIX_SEQ_I:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   160
  assumes "POSIX v1 r1" "POSIX v2 r2" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   161
  shows "POSIX (Seq v1 v2) (SEQ r1 r2)" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   162
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   163
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   164
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   165
apply(rotate_tac 2)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   166
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   167
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   168
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   169
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   170
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   171
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   172
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   173
*)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   174
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   175
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   176
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   177
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   178
lemma POSIX_ALT2:
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   179
  assumes "POSIX (Left v1) (ALT r1 r2)"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   180
  shows "POSIX v1 r1"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   181
using assms
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   182
unfolding POSIX_def
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   183
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   184
apply(drule_tac x="Left v'" in spec)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   185
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   186
apply(drule mp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   187
apply(rule Prf.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   188
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   189
apply(erule ValOrd.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   190
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   191
done
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   192
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   193
lemma POSIX2_ALT:
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   194
  assumes "POSIX2 (Left v1) (ALT r1 r2)"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   195
  shows "POSIX2 v1 r1"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   196
using assms
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   197
unfolding POSIX2_def
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   198
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   199
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   200
done
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   201
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   202
lemma POSIX_ALT:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   203
  assumes "POSIX (Left v1) (ALT r1 r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   204
  shows "POSIX v1 r1"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   205
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   206
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   207
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   208
apply(drule_tac x="Left v'" in spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   209
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   210
apply(drule mp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   211
apply(rule Prf.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   212
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   213
apply(erule ValOrd.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   214
apply(simp_all)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   215
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   216
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   217
lemma POSIX2_ALT:
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   218
  assumes "POSIX2 (Left v1) (ALT r1 r2)"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   219
  shows "POSIX2 v1 r1"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   220
using assms
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   221
apply(simp add: POSIX2_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   222
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   223
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   224
apply(simp_all)[5]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   225
apply(drule_tac x="Left v'" in spec)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   226
apply(drule mp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   227
apply(rule Prf.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   228
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   229
apply(erule ValOrd.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   230
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   231
done
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   232
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   233
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   234
lemma POSIX_ALT1a:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   235
  assumes "POSIX (Right v2) (ALT r1 r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   236
  shows "POSIX v2 r2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   237
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   238
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   239
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   240
apply(drule_tac x="Right v'" in spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   241
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   242
apply(drule mp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   243
apply(rule Prf.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   244
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   245
apply(erule ValOrd.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   246
apply(simp_all)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   247
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   248
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   249
lemma POSIX2_ALT1a:
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   250
  assumes "POSIX2 (Right v2) (ALT r1 r2)"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   251
  shows "POSIX2 v2 r2"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   252
using assms
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   253
unfolding POSIX2_def
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   254
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   255
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   256
apply(simp_all)[5]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   257
apply(drule_tac x="Right v'" in spec)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   258
apply(drule mp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   259
apply(rule Prf.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   260
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   261
apply(erule ValOrd.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   262
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   263
done
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   264
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   265
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   266
lemma POSIX_ALT1b:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   267
  assumes "POSIX (Right v2) (ALT r1 r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   268
  shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   269
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   270
apply(drule_tac POSIX_ALT1a)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   271
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   272
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   273
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   274
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   275
lemma POSIX_ALT_I1:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   276
  assumes "POSIX v1 r1" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   277
  shows "POSIX (Left v1) (ALT r1 r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   278
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   279
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   280
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   281
apply(rotate_tac 3)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   282
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   283
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   284
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   285
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   286
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   287
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   288
by simp
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   289
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   290
lemma POSIX2_ALT_I1:
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   291
  assumes "POSIX2 v1 r1" 
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   292
  shows "POSIX2 (Left v1) (ALT r1 r2)"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   293
using assms
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   294
unfolding POSIX2_def
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   295
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   296
apply(rule Prf.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   297
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   298
apply(rotate_tac 2)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   299
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   300
apply(simp_all)[5]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   301
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   302
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   303
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   304
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   305
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   306
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   307
by simp
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   308
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   309
lemma POSIX_ALT_I2:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   310
  assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   311
  shows "POSIX (Right v2) (ALT r1 r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   312
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   313
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   314
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   315
apply(rotate_tac 3)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   316
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   317
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   318
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   319
apply(rule ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   320
apply metis
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   321
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   322
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   323
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   324
section {* The ordering is reflexive *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   325
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   326
lemma ValOrd_refl:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   327
  assumes "\<turnstile> v : r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   328
  shows "v \<succ>r v"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   329
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   330
apply(induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   331
apply(auto intro: ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   332
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   333
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   334
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   335
section {* The Matcher *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   336
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   337
fun
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   338
 nullable :: "rexp \<Rightarrow> bool"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   339
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   340
  "nullable (NULL) = False"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   341
| "nullable (EMPTY) = True"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   342
| "nullable (CHAR c) = False"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   343
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   344
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   345
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   346
lemma nullable_correctness:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   347
  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   348
apply (induct r) 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   349
apply(auto simp add: Sequ_def) 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   350
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   351
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   352
fun mkeps :: "rexp \<Rightarrow> val"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   353
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   354
  "mkeps(EMPTY) = Void"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   355
| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   356
| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   357
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   358
lemma mkeps_nullable:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   359
  assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   360
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   361
apply(induct rule: nullable.induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   362
apply(auto intro: Prf.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   363
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   364
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   365
lemma mkeps_flat:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   366
  assumes "nullable(r)" shows "flat (mkeps r) = []"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   367
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   368
apply(induct rule: nullable.induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   369
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   370
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   371
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   372
text {*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   373
  The value mkeps returns is always the correct POSIX
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   374
  value.
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   375
*}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   376
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   377
lemma mkeps_POSIX2:
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   378
  assumes "nullable r"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   379
  shows "POSIX2 (mkeps r) r"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   380
using assms
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   381
apply(induct r)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   382
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   383
apply(simp add: POSIX2_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   384
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   385
lemma mkeps_POSIX:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   386
  assumes "nullable r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   387
  shows "POSIX (mkeps r) r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   388
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   389
apply(induct r)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   390
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   391
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   392
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   393
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   394
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   395
apply (metis ValOrd.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   396
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   397
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   398
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   399
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   400
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   401
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   402
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   403
apply (simp add: ValOrd.intros(2) mkeps_flat)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   404
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   405
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   406
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   407
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   408
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   409
apply (simp add: ValOrd.intros(6))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   410
apply (simp add: ValOrd.intros(3))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   411
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   412
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   413
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   414
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   415
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   416
apply (simp add: ValOrd.intros(6))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   417
apply (simp add: ValOrd.intros(3))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   418
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   419
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   420
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   421
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   422
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   423
apply (metis Prf_flat_L mkeps_flat nullable_correctness)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   424
by (simp add: ValOrd.intros(5))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   425
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   426
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   427
lemma mkeps_POSIX2:
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   428
  assumes "nullable r"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   429
  shows "POSIX2 (mkeps r) r"
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   430
using assms
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   431
apply(induct r)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   432
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   433
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   434
apply(simp add: POSIX2_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   435
apply(rule conjI)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   436
apply(rule Prf.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   437
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   438
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   439
apply(simp_all)[5]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   440
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   441
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   442
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   443
apply(simp add: POSIX2_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   444
apply(rule conjI)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   445
apply(rule Prf.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   446
apply(simp add: mkeps_nullable)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   447
apply(simp add: mkeps_nullable)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   448
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   449
apply(rotate_tac 6)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   450
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   451
apply(simp_all)[5]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   452
apply(rule ValOrd.intros(2))
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   453
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   454
apply(simp only: nullable.simps)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   455
apply(erule disjE)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   456
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   457
thm POSIX2_ALT1a
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   458
apply(rule POSIX2_ALT)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   459
apply(simp add: POSIX2_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   460
apply(rule conjI)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   461
apply(rule Prf.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   462
apply(simp add: mkeps_nullable)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   463
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   464
apply(rotate_tac 4)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   465
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   466
apply(simp_all)[5]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   467
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   468
apply(simp)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   469
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   470
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   471
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   472
section {* Derivatives *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   473
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   474
fun
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   475
 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   476
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   477
  "der c (NULL) = NULL"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   478
| "der c (EMPTY) = NULL"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   479
| "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   480
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   481
| "der c (SEQ r1 r2) = 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   482
     (if nullable r1
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   483
      then ALT (SEQ (der c r1) r2) (der c r2)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   484
      else SEQ (der c r1) r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   485
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   486
fun 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   487
 ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   488
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   489
  "ders [] r = r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   490
| "ders (c # s) r = ders s (der c r)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   491
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   492
section {* Injection function *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   493
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   494
fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   495
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   496
  "injval (CHAR d) c Void = Char d"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   497
| "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   498
| "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   499
| "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   500
| "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   501
| "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   502
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   503
section {* Projection function *}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   504
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   505
fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   506
where
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   507
  "projval (CHAR d) c _ = Void"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   508
| "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   509
| "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   510
| "projval (SEQ r1 r2) c (Seq v1 v2) = 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   511
     (if flat v1 = [] then Right(projval r2 c v2) 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   512
      else if nullable r1 then Left (Seq (projval r1 c v1) v2)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   513
                          else Seq (projval r1 c v1) v2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   514
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   515
text {*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   516
  Injection value is related to r
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   517
*}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   518
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   519
lemma v3:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   520
  assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   521
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   522
apply(induct arbitrary: v rule: der.induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   523
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   524
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   525
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   526
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   527
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   528
apply(case_tac "c = c'")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   529
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   530
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   531
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   532
apply (metis Prf.intros(5))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   533
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   534
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   535
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   536
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   537
apply (metis Prf.intros(2))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   538
apply (metis Prf.intros(3))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   539
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   540
apply(case_tac "nullable r1")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   541
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   542
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   543
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   544
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   545
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   546
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   547
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   548
apply (metis Prf.intros(1))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   549
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   550
apply (metis Prf.intros(1) mkeps_nullable)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   551
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   552
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   553
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   554
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   555
apply(rule Prf.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   556
apply(auto)[2]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   557
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   558
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   559
text {*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   560
  The string behin the injection value is an added c
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   561
*}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   562
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   563
lemma v4:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   564
  assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   565
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   566
apply(induct arbitrary: v rule: der.induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   567
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   568
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   569
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   570
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   571
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   572
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   573
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   574
apply(case_tac "c = c'")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   575
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   576
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   577
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   578
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   579
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   580
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   581
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   582
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   583
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   584
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   585
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   586
apply(case_tac "nullable r1")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   587
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   588
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   589
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   590
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   591
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   592
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   593
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   594
apply (metis mkeps_flat)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   595
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   596
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   597
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   598
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   599
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   600
text {*
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   601
  Injection followed by projection is the identity.
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   602
*}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   603
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   604
lemma proj_inj_id:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   605
  assumes "\<turnstile> v : der c r" 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   606
  shows "projval r c (injval r c v) = v"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   607
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   608
apply(induct r arbitrary: c v rule: rexp.induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   609
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   610
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   611
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   612
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   613
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   614
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   615
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   616
apply(case_tac "c = char")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   617
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   618
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   619
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   620
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   621
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   622
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   623
defer
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   624
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   625
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   626
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   627
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   628
apply(case_tac "nullable rexp1")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   629
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   630
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   631
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   632
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   633
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   634
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   635
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   636
apply (metis list.distinct(1) v4)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   637
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   638
apply (metis mkeps_flat)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   639
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   640
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   641
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   642
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   643
apply(simp add: v4)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   644
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   645
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   646
lemma "\<exists>v. POSIX v r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   647
apply(induct r)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   648
apply(rule exI)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   649
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   650
apply (metis (full_types) Prf_flat_L der.simps(1) der.simps(2) der.simps(3) flat.simps(1) nullable.simps(1) nullable_correctness proj_inj_id projval.simps(1) v3 v4)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   651
apply(rule_tac x = "Void" in exI)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   652
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   653
apply (metis POSIX_def flat.simps(1) mkeps.simps(1) mkeps_POSIX nullable.simps(2))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   654
apply(rule_tac x = "Char char" in exI)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   655
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   656
apply(auto) [1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   657
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   658
apply(simp_all) [5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   659
apply (metis ValOrd.intros(8))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   660
defer
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   661
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   662
apply (metis POSIX_ALT_I1)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   663
(* maybe it is too early to instantiate this existential quantifier *)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   664
(* potentially this is the wrong POSIX value *)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   665
apply(rule_tac x = "Seq v va" in exI )
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   666
apply(simp (no_asm) add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   667
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   668
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   669
apply(simp_all)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   670
apply(case_tac "v \<succ>r1a v1")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   671
apply (metis ValOrd.intros(2))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   672
apply(simp add: POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   673
apply(case_tac "flat v = flat v1")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   674
apply(auto)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   675
apply(simp only: append_eq_append_conv2)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   676
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   677
thm append_eq_append_conv2
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   678
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   679
text {* 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   680
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   681
  HERE: Crucial lemma that does not go through in the sequence case. 
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   682
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   683
*}
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   684
lemma v5:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   685
  assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   686
  shows "POSIX (injval r c v) r"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   687
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   688
apply(induct arbitrary: v rule: der.induct)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   689
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   690
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   691
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   692
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   693
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   694
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   695
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   696
apply(case_tac "c = c'")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   697
apply(auto simp add: POSIX_def)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   698
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   699
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   700
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   701
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   702
using ValOrd.simps apply blast
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   703
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   704
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   705
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   706
(* base cases done *)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   707
(* ALT case *)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   708
apply(erule Prf.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   709
apply(simp_all)[5]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   710
using POSIX_ALT POSIX_ALT_I1 apply blast
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   711
apply(clarify)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   712
apply(subgoal_tac "POSIX v2 (der c r2)")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   713
prefer 2
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   714
apply(auto simp add: POSIX_def)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   715
apply (metis POSIX_ALT1a POSIX_def flat.simps(4))
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   716
apply(rotate_tac 1)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   717
apply(drule_tac x="v2" in meta_spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   718
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   719
apply(subgoal_tac "\<turnstile> Right (injval r2 c v2) : (ALT r1 r2)")
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   720
prefer 2
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   721
apply (metis Prf.intros(3) v3)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   722
apply(rule ccontr)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   723
apply(auto simp add: POSIX_def)[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   724
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   725
apply(rule allI)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   726
apply(rule impI)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   727
apply(erule conjE)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   728
thm POSIX_ALT_I2
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   729
apply(frule POSIX_ALT1a)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   730
apply(drule POSIX_ALT1b)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   731
apply(rule POSIX_ALT_I2)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   732
apply auto[1]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   733
apply(subst v4)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   734
apply(auto)[2]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   735
apply(rotate_tac 1)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   736
apply(drule_tac x="v2" in meta_spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   737
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   738
apply(subst (asm) (4) POSIX_def)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   739
apply(subst (asm) v4)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   740
apply(auto)[2]
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
   741
(* stuck in the ALT case *)