thys/Re1.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Mon, 25 May 2015 16:09:23 +0100
changeset 77 4b4c677501e7
parent 76 39cef7b9212a
child 78 279d0bc48308
permissions -rw-r--r--
proved some basic properties (totality and trichonomity) for the orderings
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56
5bc72d6d633d updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 55
diff changeset
     1
   
11
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
     2
theory Re1
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
  imports "Main" 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
begin
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
section {* Sequential Composition of Sets *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
definition
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
  Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
where 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
  "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
text {* Two Simple Properties about Sequential Composition *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
lemma seq_empty [simp]:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
  shows "A ;; {[]} = A"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
  and   "{[]} ;; A = A"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
by (simp_all add: Sequ_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
lemma seq_null [simp]:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
  shows "A ;; {} = {}"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
  and   "{} ;; A = {}"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
by (simp_all add: Sequ_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
section {* Regular Expressions *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
datatype rexp =
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
  NULL
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
| EMPTY
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
| CHAR char
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
| SEQ rexp rexp
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
| ALT rexp rexp
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    34
fun SEQS :: "rexp \<Rightarrow> rexp list \<Rightarrow> rexp"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    35
where
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    36
  "SEQS r [] = r"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    37
| "SEQS r (r'#rs) = SEQ r (SEQS r' rs)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    38
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
section {* Semantics of Regular Expressions *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
fun
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
  L :: "rexp \<Rightarrow> string set"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
  "L (NULL) = {}"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
| "L (EMPTY) = {[]}"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
| "L (CHAR c) = {[c]}"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
| "L (SEQ r1 r2) = (L r1) ;; (L r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
| "L (ALT r1 r2) = (L r1) \<union> (L r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    50
fun
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    51
 nullable :: "rexp \<Rightarrow> bool"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    52
where
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    53
  "nullable (NULL) = False"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    54
| "nullable (EMPTY) = True"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    55
| "nullable (CHAR c) = False"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    56
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    57
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
70
fahad
parents: 68
diff changeset
    58
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    59
fun
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    60
 nullable_left :: "rexp \<Rightarrow> bool"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    61
where
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    62
  "nullable_left (NULL) = False"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    63
| "nullable_left (EMPTY) = True"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    64
| "nullable_left (CHAR c) = False"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    65
| "nullable_left (ALT r1 r2) = (nullable_left r1 \<or> nullable_left r2)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    66
| "nullable_left (SEQ r1 r2) = nullable_left r1"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    67
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    68
lemma nullable_left:
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    69
  "nullable r \<Longrightarrow> nullable_left r"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    70
apply(induct r)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    71
apply(auto)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    72
done
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    73
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    74
54
45274393f28c no changes
fahad
parents: 53
diff changeset
    75
value "L(CHAR c)"
45274393f28c no changes
fahad
parents: 53
diff changeset
    76
value "L(SEQ(CHAR c)(CHAR b))"
45274393f28c no changes
fahad
parents: 53
diff changeset
    77
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    78
lemma nullable_correctness:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    79
  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    80
apply (induct r) 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    81
apply(auto simp add: Sequ_def) 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    82
done
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
section {* Values *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
datatype val = 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
  Void
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
| Char char
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
| Seq val val
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
| Right val
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
| Left val
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    93
fun Seqs :: "val \<Rightarrow> val list \<Rightarrow> val"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    94
where
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    95
  "Seqs v [] = v"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    96
| "Seqs v (v'#vs) = Seq v (Seqs v' vs)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    97
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    98
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
section {* The string behind a value *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
fun flat :: "val \<Rightarrow> string"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
  "flat(Void) = []"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
| "flat(Char c) = [c]"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
| "flat(Left v) = flat(v)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
| "flat(Right v) = flat(v)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
| "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   109
fun flat_left :: "val \<Rightarrow> string"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   110
where
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   111
  "flat_left(Void) = []"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   112
| "flat_left(Char c) = [c]"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   113
| "flat_left(Left v) = flat_left(v)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   114
| "flat_left(Right v) = flat_left(v)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   115
| "flat_left(Seq v1 v2) = flat_left(v1)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   116
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   117
fun flat_right :: "val \<Rightarrow> string"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   118
where
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   119
  "flat_right(Void) = []"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   120
| "flat_right(Char c) = [c]"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   121
| "flat_right(Left v) = flat(v)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   122
| "flat_right(Right v) = flat(v)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   123
| "flat_right(Seq v1 v2) = flat(v2)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   124
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   125
fun head :: "val \<Rightarrow> string"
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   126
where
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   127
  "head(Void) = []"
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   128
| "head(Char c) = [c]"
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   129
| "head(Left v) = head(v)"
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   130
| "head(Right v) = head(v)"
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   131
| "head(Seq v1 v2) = head v1"
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
   132
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   133
fun flats :: "val \<Rightarrow> string list"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   134
where
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   135
  "flats(Void) = [[]]"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   136
| "flats(Char c) = [[c]]"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   137
| "flats(Left v) = flats(v)"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   138
| "flats(Right v) = flats(v)"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   139
| "flats(Seq v1 v2) = (flats v1) @ (flats v2)"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
70
fahad
parents: 68
diff changeset
   141
value "flats(Seq(Char c)(Char b))"
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   142
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   143
section {* Relation between values and regular expressions *}
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   144
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   145
inductive Prfs :: "string \<Rightarrow> val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile>_ _ : _" [100, 100, 100] 100)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   146
where
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   147
 "\<lbrakk>\<Turnstile>s1 v1 : r1; \<Turnstile>s2 v2 : r2\<rbrakk> \<Longrightarrow> \<Turnstile>(s1 @ s2) (Seq v1 v2) : SEQ r1 r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   148
| "\<Turnstile>s v1 : r1 \<Longrightarrow> \<Turnstile>s (Left v1) : ALT r1 r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   149
| "\<Turnstile>s v2 : r2 \<Longrightarrow> \<Turnstile>s (Right v2) : ALT r1 r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   150
| "\<Turnstile>[] Void : EMPTY"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   151
| "\<Turnstile>[c] (Char c) : CHAR c"
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   152
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   153
lemma Prfs_flat:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   154
  "\<Turnstile>s v : r \<Longrightarrow> flat v = s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   155
apply(induct s v r rule: Prfs.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   156
apply(auto)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   157
done
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   158
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   159
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   160
inductive Prfn :: "nat \<Rightarrow> val \<Rightarrow> rexp \<Rightarrow> bool" ("\<TTurnstile>_ _ : _" [100, 100, 100] 100)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   161
where
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   162
 "\<lbrakk>\<TTurnstile>n1 v1 : r1; \<TTurnstile>n2 v2 : r2\<rbrakk> \<Longrightarrow> \<TTurnstile>(n1 + n2) (Seq v1 v2) : SEQ r1 r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   163
| "\<TTurnstile>n v1 : r1 \<Longrightarrow> \<TTurnstile>n (Left v1) : ALT r1 r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   164
| "\<TTurnstile>n v2 : r2 \<Longrightarrow> \<TTurnstile>n (Right v2) : ALT r1 r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   165
| "\<TTurnstile>0 Void : EMPTY"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   166
| "\<TTurnstile>1 (Char c) : CHAR c"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   167
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   168
lemma Prfn_flat:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   169
  "\<TTurnstile>n v : r \<Longrightarrow> length(flat v) = n"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   170
apply(induct rule: Prfn.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   171
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   172
done
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
 "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
| "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
| "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
| "\<turnstile> Void : EMPTY"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
| "\<turnstile> Char c : CHAR c"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   182
lemma Prf_Prfn:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   183
  shows "\<turnstile> v : r \<Longrightarrow> \<TTurnstile>(length (flat v)) v : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   184
apply(induct v r rule: Prf.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   185
apply(auto intro: Prfn.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   186
by (metis One_nat_def Prfn.intros(5))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   187
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   188
lemma Prfn_Prf:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   189
  shows "\<TTurnstile>n v : r \<Longrightarrow> \<turnstile> v : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   190
apply(induct n v r rule: Prfn.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   191
apply(auto intro: Prf.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   192
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   193
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   194
lemma Prf_Prfs:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   195
  shows "\<turnstile> v : r \<Longrightarrow> \<Turnstile>(flat v) v : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   196
apply(induct v r rule: Prf.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   197
apply(auto intro: Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   198
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   199
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   200
lemma Prfs_Prf:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   201
  shows "\<Turnstile>s v : r \<Longrightarrow> \<turnstile> v : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   202
apply(induct s v r rule: Prfs.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   203
apply(auto intro: Prf.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   204
done
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   205
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   206
fun mkeps :: "rexp \<Rightarrow> val"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
where
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   208
  "mkeps(EMPTY) = Void"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   209
| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   210
| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   212
lemma mkeps_nullable:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   213
  assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   214
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   215
apply(induct rule: nullable.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   216
apply(auto intro: Prf.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   217
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   218
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   219
lemma mkeps_nullable_n:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   220
  assumes "nullable(r)" shows "\<TTurnstile>0 (mkeps r) : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   221
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   222
apply(induct rule: nullable.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   223
apply(auto intro: Prfn.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   224
apply(drule Prfn.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   225
apply(assumption)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   226
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   227
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   228
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   229
lemma mkeps_nullable_s:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   230
  assumes "nullable(r)" shows "\<Turnstile>[] (mkeps r) : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   231
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   232
apply(induct rule: nullable.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   233
apply(auto intro: Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   234
apply(drule Prfs.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   235
apply(assumption)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   236
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   237
done
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   238
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   239
lemma mkeps_flat:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   240
  assumes "nullable(r)" shows "flat (mkeps r) = []"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   241
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   242
apply(induct rule: nullable.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   243
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   244
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   245
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   246
lemma mkeps_flat_left:
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   247
  assumes "nullable(r)" shows "flat_left (mkeps r) = []"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   248
using assms
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   249
apply(induct rule: nullable.induct)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   250
apply(auto)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   251
done
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   252
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   253
text {*
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   254
  The value mkeps returns is always the correct POSIX
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   255
  value.
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   256
*}
70
fahad
parents: 68
diff changeset
   257
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
lemma Prf_flat_L:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
  assumes "\<turnstile> v : r" shows "flat v \<in> L r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
using assms
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   261
apply(induct v r rule: Prf.induct)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
apply(auto simp add: Sequ_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
lemma L_flat_Prf:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
  "L(r) = {flat v | v. \<turnstile> v : r}"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
apply(induct r)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
apply(auto dest: Prf_flat_L simp add: Sequ_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
apply (metis Prf.intros(4) flat.simps(1))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
apply (metis Prf.intros(5) flat.simps(2))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
apply (metis Prf.intros(1) flat.simps(5))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
apply (metis Prf.intros(2) flat.simps(3))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
apply (metis Prf.intros(3) flat.simps(4))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
70
fahad
parents: 68
diff changeset
   278
definition prefix :: "string \<Rightarrow> string \<Rightarrow> bool" ("_ \<sqsubset> _" [100, 100] 100)
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   279
where
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   280
  "s1 \<sqsubset> s2 \<equiv> \<exists>s3. s1 @ s3 = s2"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   281
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
section {* Ordering of values *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
where
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   286
  "v2 \<succ>r2 v2' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1 v2')" 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   287
| "\<lbrakk>v1 \<succ>r1 v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
| "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
| "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
| "Void \<succ>EMPTY Void"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
| "(Char c) \<succ>(CHAR c) (Char c)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   295
inductive ValOrd2 :: "val \<Rightarrow> val \<Rightarrow> bool" ("_ 2\<succ> _" [100, 100] 100)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   296
where
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   297
  "v2 2\<succ> v2' \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1 v2')" 
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   298
| "\<lbrakk>v1 2\<succ> v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1' v2')" 
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   299
| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) 2\<succ> (Right v2)"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   300
| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) 2\<succ> (Left v1)"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   301
| "v2 2\<succ> v2' \<Longrightarrow> (Right v2) 2\<succ> (Right v2')"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   302
| "v1 2\<succ> v1' \<Longrightarrow> (Left v1) 2\<succ> (Left v1')"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   303
| "Void 2\<succ> Void"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   304
| "(Char c) 2\<succ> (Char c)"
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   305
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   306
lemma Ord1:
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   307
  "v1 \<succ>r v2 \<Longrightarrow> v1 2\<succ> v2"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   308
apply(induct rule: ValOrd.induct)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   309
apply(auto intro: ValOrd2.intros)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   310
done
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   311
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   312
lemma Ord2:
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   313
  "v1 2\<succ> v2 \<Longrightarrow> \<exists>r. v1 \<succ>r v2"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   314
apply(induct v1 v2 rule: ValOrd2.induct)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   315
apply(auto intro: ValOrd.intros)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   316
done
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   317
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   318
lemma Ord3:
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   319
  "\<lbrakk>v1 2\<succ> v2; \<turnstile> v1 : r\<rbrakk> \<Longrightarrow> v1 \<succ>r v2"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   320
apply(induct v1 v2 arbitrary: r rule: ValOrd2.induct)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   321
apply(auto intro: ValOrd.intros elim: Prf.cases)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   322
done
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   323
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   324
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   325
lemma ValOrd_refl:
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   326
  assumes "\<turnstile> v : r"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   327
  shows "v \<succ>r v"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
using assms
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   329
apply(induct)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   330
apply(auto intro: ValOrd.intros)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
done
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   332
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   333
lemma 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   334
  "flat Void = []"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   335
  "flat (Seq Void Void) = []"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   336
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   337
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   338
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   339
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   340
lemma ValOrd_total:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   341
  shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r\<rbrakk>  \<Longrightarrow> v1 \<succ>r v2 \<or> v2 \<succ>r v1"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   342
apply(induct r arbitrary: v1 v2)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   343
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   344
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   345
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   346
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   347
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   348
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   349
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   350
apply (metis ValOrd.intros(7))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   351
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   352
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   353
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   354
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   355
apply (metis ValOrd.intros(8))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   356
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   357
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   358
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   359
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   360
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   361
apply(case_tac "v1a = v1b")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   362
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   363
apply(rule ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   364
apply (metis ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   365
apply(rule ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   366
apply(auto)[2]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   367
apply(erule contrapos_np)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   368
apply(rule ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   369
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   370
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   371
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   372
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   373
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   374
apply (metis Ord1 Ord3 Prf.intros(2) ValOrd2.intros(6))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   375
apply(rule ValOrd.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   376
apply(erule contrapos_np)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   377
apply(rule ValOrd.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   378
apply (metis le_eq_less_or_eq neq_iff)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   379
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   380
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   381
apply(rule ValOrd.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   382
apply(erule contrapos_np)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   383
apply(rule ValOrd.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   384
apply (metis le_eq_less_or_eq neq_iff)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   385
apply(rule ValOrd.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   386
apply(erule contrapos_np)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   387
apply(rule ValOrd.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   388
by metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   389
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   390
lemma ValOrd_anti:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   391
  shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r; v1 \<succ>r v2; v2 \<succ>r v1\<rbrakk> \<Longrightarrow> v1 = v2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   392
apply(induct r arbitrary: v1 v2)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   393
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   394
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   395
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   396
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   397
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   398
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   399
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   400
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   401
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   402
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   403
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   404
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   405
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   406
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   407
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   408
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   409
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   410
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   411
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   412
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   413
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   414
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   415
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   416
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   417
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   418
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   419
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   420
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   421
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   422
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   423
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   424
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   425
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   426
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   427
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   428
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   429
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   430
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   431
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   432
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   433
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   434
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   435
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   436
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   437
lemma ValOrd_max:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   438
  shows "\<exists>v. \<forall>v'. (v' \<succ>r v \<longrightarrow> v = v')" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   439
apply(induct r)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   440
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   441
apply(rule exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   442
apply(rule allI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   443
apply(rule impI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   444
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   445
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   446
apply(rule exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   447
apply(rule allI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   448
apply(rule impI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   449
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   450
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   451
apply(rule exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   452
apply(rule allI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   453
apply(rule impI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   454
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   455
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   456
apply(rule exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   457
apply(rule allI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   458
apply(rule impI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   459
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   460
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   461
apply(rule exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   462
apply(rule allI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   463
apply(rule impI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   464
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   465
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   466
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   467
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   468
lemma ValOrd_max:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   469
  assumes "L r \<noteq> {}"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   470
  shows "\<exists>v. \<turnstile> v : r \<and> (\<forall>v'. ((\<turnstile> v' : r \<and> v' \<succ>r v) \<longrightarrow> v = v'))" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   471
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   472
apply(induct r)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   473
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   474
apply(rule exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   475
apply(rule conjI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   476
apply(rule Prf.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   477
apply(rule allI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   478
apply(rule impI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   479
apply(erule conjE)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   480
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   481
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   482
apply(rule exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   483
apply(rule conjI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   484
apply(rule Prf.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   485
apply(rule allI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   486
apply(rule impI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   487
apply(erule conjE)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   488
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   489
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   490
apply(auto simp add: Sequ_def)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   491
apply(drule meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   492
apply (metis empty_iff)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   493
apply(drule meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   494
apply (metis empty_iff)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   495
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   496
apply(drule_tac x="Seq v va" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   497
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   498
apply (metis Prf.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   499
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   500
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   501
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   502
apply(rotate_tac 6)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   503
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   504
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   505
apply(rotate_tac 6)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   506
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   507
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   508
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   509
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   510
apply(drule meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   511
apply (metis empty_iff)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   512
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   513
apply(drule_tac x="Left v" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   514
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   515
apply (metis Prf.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   516
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   517
apply(rotate_tac 4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   518
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   519
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   520
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   521
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   522
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   523
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   524
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   525
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   526
oops
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   527
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   528
lemma ValOrd_max2:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   529
  shows "\<exists>v. \<turnstile> v : r \<and> (\<forall>v'. v \<succ>r v')" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   530
using ValOrd_max[where r="r"]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   531
apply -
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   532
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   533
apply(rule_tac x="v" in exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   534
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   535
oops
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   536
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   537
lemma ValOrd_trans:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   538
  assumes "v1 \<succ>r v2" "v2 \<succ>r v3" "\<turnstile> v1 : r" "\<turnstile> v2 : r" "\<turnstile> v3 : r" "flat v1 = flat v2" "flat v2 = flat v3"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   539
  shows "v1 \<succ>r v3"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   540
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   541
apply(induct r arbitrary: v1 v2 v3)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   542
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   543
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   544
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   545
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   546
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   547
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   548
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   549
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   550
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   551
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   552
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   553
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   554
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   555
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   556
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   557
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   558
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   559
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   560
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   561
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   562
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   563
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   564
apply (metis ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   565
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   566
apply (metis ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   567
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   568
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   569
apply (metis ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   570
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   571
apply(case_tac "v1d = v1'a")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   572
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   573
apply (metis ValOrd_anti)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   574
apply (rule ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   575
prefer 2
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   576
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   577
prefer 2
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   578
oops
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   579
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   580
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   581
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   582
section {* Posix definition *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   583
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   584
definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   585
where
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   586
  "POSIX v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v'))"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   587
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   588
definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   589
where
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   590
  "POSIX2 v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v 2\<succ> v'))"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   591
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   592
lemma "POSIX v r = POSIX2 v r"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   593
unfolding POSIX_def POSIX2_def
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   594
apply(auto)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   595
apply(rule Ord1)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   596
apply(auto)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   597
apply(rule Ord3)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   598
apply(auto)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   599
done
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   600
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   601
definition POSIXs :: "val \<Rightarrow> rexp \<Rightarrow> string \<Rightarrow> bool" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   602
where
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   603
  "POSIXs v r s \<equiv> (\<Turnstile>s v : r \<and> (\<forall>v'. (\<Turnstile>s v' : r \<longrightarrow> v 2\<succ> v')))"
11
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   604
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   605
definition POSIXn :: "val \<Rightarrow> rexp \<Rightarrow> nat \<Rightarrow> bool" 
11
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   606
where
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   607
  "POSIXn v r n \<equiv> (\<TTurnstile>n v : r \<and> (\<forall>v'. (\<TTurnstile>n v' : r \<longrightarrow> v 2\<succ> v')))"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   608
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   609
lemma "POSIXn v r (length (flat v)) \<Longrightarrow> POSIX2 v r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   610
unfolding POSIXn_def POSIX2_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   611
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   612
apply (metis Prfn_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   613
by (metis Prf_Prfn)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   614
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   615
lemma Prfs_POSIX:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   616
  "POSIXs v r s \<Longrightarrow> \<Turnstile>s v: r \<and> flat v = s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   617
apply(simp add: POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   618
by (metis Prfs_flat)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   619
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   620
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   621
lemma "POSIXs v r (flat v) =  POSIX2 v r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   622
unfolding POSIXs_def POSIX2_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   623
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   624
apply (metis Prfs_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   625
apply (metis Prf_Prfs)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   626
apply (metis Prf_Prfs)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   627
by (metis Prfs_Prf Prfs_flat)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   628
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   629
lemma 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   630
  assumes "POSIX v1 r1" "\<turnstile> v1' : r1"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   631
  shows "Seq v1 v2 \<succ>(SEQ r1 r2) Seq v1' v2'"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   632
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   633
apply(rule_tac ValOrd.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   634
apply(simp add: POSIX_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   635
oops
11
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   636
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   637
lemma
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   638
"L r \<noteq> {} \<Longrightarrow> \<exists>v. POSIXs v r (flat v)"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   639
apply(induct r arbitrary: )
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   640
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   641
apply(rule_tac x="Void" in exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   642
apply(simp add: POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   643
apply(rule conjI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   644
apply (metis Prfs.intros(4))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   645
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   646
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   647
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   648
apply (metis ValOrd2.intros(7))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   649
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   650
apply(rule_tac x="Char char" in exI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   651
apply(simp add: POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   652
apply(rule conjI)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   653
apply (metis Prfs.intros(5))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   654
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   655
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   656
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   657
apply (metis ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   658
apply(auto simp add: Sequ_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   659
apply(drule meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   660
apply (metis empty_iff)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   661
apply(drule meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   662
apply (metis empty_iff)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   663
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   664
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   665
section {* POSIX for some constructors *}
11
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   666
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   667
lemma POSIX_SEQ1:
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   668
  assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   669
  shows "POSIX v1 r1"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   673
apply(drule_tac x="Seq v' v2" in spec)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   674
apply(simp)
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   675
apply(erule impE)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   676
apply(rule Prf.intros)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   677
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   678
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   679
apply(erule ValOrd.cases)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   680
apply(simp_all)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   681
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   682
by (metis ValOrd_refl)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   683
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   684
lemma POSIXn_SEQ1:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   685
  assumes "POSIXn (Seq v1 v2) (SEQ r1 r2) (n1 + n2)" "\<TTurnstile>n1 v1 : r1" "\<TTurnstile>n2 v2 : r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   686
  shows "POSIXn v1 r1 n1"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   687
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   688
unfolding POSIXn_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   689
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   690
apply(drule_tac x="Seq v' v2" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   691
apply(erule impE)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   692
apply(rule Prfn.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   693
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   694
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   695
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   696
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   697
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   698
by (metis Ord1 Prfn_Prf ValOrd_refl)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   699
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   700
lemma POSIXs_SEQ1:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   701
  assumes "POSIXs (Seq v1 v2) (SEQ r1 r2) (s1 @ s2)" "\<Turnstile>s1 v1 : r1" "\<Turnstile>s2 v2 : r2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   702
  shows "POSIXs v1 r1 s1"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   703
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   704
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   705
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   706
apply(drule_tac x="Seq v' v2" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   707
apply(erule impE)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   708
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   709
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   710
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   711
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   712
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   713
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   714
by (metis Ord1 Prfs_Prf ValOrd_refl)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   715
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   716
lemma POSIX_SEQ2:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   717
  assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2" 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   718
  shows "POSIX v2 r2"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   719
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   720
unfolding POSIX_def
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   721
apply(auto)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   722
apply(drule_tac x="Seq v1 v'" in spec)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   723
apply(simp)
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   724
apply(erule impE)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   725
apply(rule Prf.intros)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   726
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   727
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   728
apply(erule ValOrd.cases)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   729
apply(simp_all)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   730
done
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   731
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   732
lemma POSIXn_SEQ2:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   733
  assumes "POSIXn (Seq v1 v2) (SEQ r1 r2) (n1 + n2)" "\<TTurnstile>n1 v1 : r1" "\<TTurnstile>n2 v2 : r2" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   734
  shows "POSIXn v2 r2 n2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   735
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   736
unfolding POSIXn_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   737
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   738
apply(drule_tac x="Seq v1 v'" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   739
apply(erule impE)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   740
apply(rule Prfn.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   741
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   742
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   743
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   744
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   745
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   746
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   747
lemma POSIXs_SEQ2:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   748
  assumes "POSIXs (Seq v1 v2) (SEQ r1 r2) (s1 @ s2)" "\<Turnstile>s1 v1 : r1" "\<Turnstile>s2 v2 : r2" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   749
  shows "POSIXs v2 r2 s2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   750
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   751
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   752
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   753
apply(drule_tac x="Seq v1 v'" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   754
apply(erule impE)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   755
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   756
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   757
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   758
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   759
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   760
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   761
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   762
lemma POSIX_ALT2:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   763
  assumes "POSIX (Left v1) (ALT r1 r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   764
  shows "POSIX v1 r1"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   765
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   766
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   767
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   768
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   769
apply(simp_all)[5]
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   770
apply(drule_tac x="Left v'" in spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   771
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   772
apply(drule mp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   773
apply(rule Prf.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   774
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   775
apply(erule ValOrd.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   776
apply(simp_all)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   777
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   778
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   779
lemma POSIXn_ALT2:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   780
  assumes "POSIXn (Left v1) (ALT r1 r2) n"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   781
  shows "POSIXn v1 r1 n"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   782
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   783
unfolding POSIXn_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   784
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   785
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   786
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   787
apply(drule_tac x="Left v'" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   788
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   789
apply(rule Prfn.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   790
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   791
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   792
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   793
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   794
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   795
lemma POSIXs_ALT2:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   796
  assumes "POSIXs (Left v1) (ALT r1 r2) s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   797
  shows "POSIXs v1 r1 s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   798
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   799
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   800
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   801
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   802
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   803
apply(drule_tac x="Left v'" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   804
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   805
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   806
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   807
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   808
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   809
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   810
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   811
lemma POSIX_ALT1a:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   812
  assumes "POSIX (Right v2) (ALT r1 r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   813
  shows "POSIX v2 r2"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   814
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   815
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   816
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   817
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   818
apply(simp_all)[5]
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   819
apply(drule_tac x="Right v'" in spec)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   820
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   821
apply(drule mp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   822
apply(rule Prf.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   823
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   824
apply(erule ValOrd.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   825
apply(simp_all)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   826
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   827
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   828
lemma POSIXn_ALT1a:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   829
  assumes "POSIXn (Right v2) (ALT r1 r2) n"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   830
  shows "POSIXn v2 r2 n"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   831
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   832
unfolding POSIXn_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   833
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   834
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   835
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   836
apply(drule_tac x="Right v'" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   837
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   838
apply(rule Prfn.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   839
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   840
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   841
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   842
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   843
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   844
lemma POSIXs_ALT1a:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   845
  assumes "POSIXs (Right v2) (ALT r1 r2) s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   846
  shows "POSIXs v2 r2 s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   847
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   848
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   849
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   850
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   851
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   852
apply(drule_tac x="Right v'" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   853
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   854
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   855
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   856
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   857
apply(simp_all)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   858
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   859
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   860
lemma POSIX_ALT1b:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   861
  assumes "POSIX (Right v2) (ALT r1 r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   862
  shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   863
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   864
apply(drule_tac POSIX_ALT1a)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   865
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   866
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   867
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   868
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   869
lemma POSIXn_ALT1b:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   870
  assumes "POSIXn (Right v2) (ALT r1 r2) n"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   871
  shows "(\<forall>v'. (\<TTurnstile>n v' : r2 \<longrightarrow> v2 2\<succ> v'))"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   872
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   873
apply(drule_tac POSIXn_ALT1a)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   874
unfolding POSIXn_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   875
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   876
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   877
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   878
lemma POSIXs_ALT1b:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   879
  assumes "POSIXs (Right v2) (ALT r1 r2) s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   880
  shows "(\<forall>v'. (\<Turnstile>s v' : r2 \<longrightarrow> v2 2\<succ> v'))"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   881
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   882
apply(drule_tac POSIXs_ALT1a)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   883
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   884
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   885
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   886
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   887
lemma POSIX_ALT_I1:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   888
  assumes "POSIX v1 r1" 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   889
  shows "POSIX (Left v1) (ALT r1 r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   890
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   891
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   892
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   893
apply (metis Prf.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   894
apply(rotate_tac 2)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   895
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   896
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   897
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   898
apply(rule ValOrd.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   899
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   900
apply(rule ValOrd.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   901
by simp
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   902
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   903
lemma POSIXn_ALT_I1:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   904
  assumes "POSIXn v1 r1 n" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   905
  shows "POSIXn (Left v1) (ALT r1 r2) n"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   906
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   907
unfolding POSIXn_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   908
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   909
apply (metis Prfn.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   910
apply(rotate_tac 2)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   911
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   912
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   913
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   914
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   915
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   916
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   917
by (metis Prfn_flat order_refl)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   918
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   919
lemma POSIXs_ALT_I1:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   920
  assumes "POSIXs v1 r1 s" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   921
  shows "POSIXs (Left v1) (ALT r1 r2) s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   922
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   923
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   924
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   925
apply (metis Prfs.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   926
apply(rotate_tac 2)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   927
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   928
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   929
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   930
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   931
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   932
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   933
by (metis Prfs_flat order_refl)
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   934
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   935
lemma POSIX_ALT_I2:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   936
  assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   937
  shows "POSIX (Right v2) (ALT r1 r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   938
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   939
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   940
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   941
apply (metis Prf.intros)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   942
apply(rotate_tac 3)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   943
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   944
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   945
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   946
apply(rule ValOrd.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   947
apply metis
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   948
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   949
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   950
lemma POSIXs_ALT_I2:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   951
  assumes "POSIXs v2 r2 s" "\<forall>s' v'. \<Turnstile>s' v' : r1 \<longrightarrow> length s > length s'"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   952
  shows "POSIXs (Right v2) (ALT r1 r2) s"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   953
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   954
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   955
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   956
apply (metis Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   957
apply(rotate_tac 3)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   958
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   959
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   960
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   961
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   962
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   963
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   964
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   965
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   966
lemma 
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   967
  "\<lbrakk>POSIX (mkeps r2) r2; nullable r2; \<not> nullable r1\<rbrakk>
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   968
   \<Longrightarrow> POSIX (Right (mkeps r2)) (ALT r1 r2)" 
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   969
apply(auto simp add: POSIX_def)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   970
apply(rule Prf.intros(3))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   971
apply(auto)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   972
apply(rotate_tac 3)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   973
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   974
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   975
apply(simp add: mkeps_flat)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   976
apply(auto)[1]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   977
apply (metis Prf_flat_L nullable_correctness)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   978
apply(rule ValOrd.intros)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   979
apply(auto)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   980
done
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   981
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   982
lemma mkeps_POSIX:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   983
  assumes "nullable r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   984
  shows "POSIX (mkeps r) r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   985
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   986
apply(induct r)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   987
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   988
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   989
apply(auto)[1]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   990
apply (metis Prf.intros(4))
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   991
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   992
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   993
apply (metis ValOrd.intros)
66
eb97e8361211 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   994
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   995
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   996
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   997
apply(auto)[1]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   998
apply (metis mkeps.simps(2) mkeps_nullable nullable.simps(5))
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   999
apply(rotate_tac 6)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1000
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1001
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1002
apply (simp add: mkeps_flat)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1003
apply(case_tac "mkeps r1a = v1")
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1004
apply(simp)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1005
apply (metis ValOrd.intros(1))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1006
apply (rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1007
apply metis
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1008
apply(simp)
71
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1009
(* ALT case *)
72
9128b9440e93 updated R1 and notes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
  1010
thm mkeps.simps
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1011
apply(simp)
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1012
apply(erule disjE)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1013
apply(simp)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1014
apply (metis POSIX_ALT_I1)
71
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1015
(* *)
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1016
apply(auto)[1]
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1017
thm  POSIX_ALT_I1
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1018
apply (metis POSIX_ALT_I1)
71
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1019
apply(simp (no_asm) add: POSIX_def)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1020
apply(auto)[1]
71
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1021
apply(rule Prf.intros(3))
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1022
apply(simp only: POSIX_def)
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1023
apply(rotate_tac 4)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1024
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1025
apply(simp_all)[5]
72
9128b9440e93 updated R1 and notes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
  1026
thm mkeps_flat
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1027
apply(simp add: mkeps_flat)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1028
apply(auto)[1]
71
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1029
thm Prf_flat_L nullable_correctness
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1030
apply (metis Prf_flat_L nullable_correctness)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1031
apply(rule ValOrd.intros)
71
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1032
apply(subst (asm) POSIX_def)
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1033
apply(clarify)
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1034
apply(drule_tac x="v2" in spec)
72
9128b9440e93 updated R1 and notes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
  1035
by simp
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1036
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1037
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1038
section {* Derivatives *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1039
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1040
fun
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1041
 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1042
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1043
  "der c (NULL) = NULL"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1044
| "der c (EMPTY) = NULL"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1045
| "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1046
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1047
| "der c (SEQ r1 r2) = 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1048
     (if nullable r1
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1049
      then ALT (SEQ (der c r1) r2) (der c r2)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1050
      else SEQ (der c r1) r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1051
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1052
fun 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1053
 ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1054
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1055
  "ders [] r = r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1056
| "ders (c # s) r = ders s (der c r)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1057
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1058
section {* Injection function *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1059
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1060
fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1061
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1062
  "injval (CHAR d) c Void = Char d"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1063
| "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1064
| "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1065
| "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1066
| "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1067
| "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1068
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1069
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1070
section {* Projection function *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1071
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1072
fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1073
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1074
  "projval (CHAR d) c _ = Void"
71
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1075
| "projval (ALT r1 r2) c (Left v1) = Left (projval r1 c v1)"
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1076
| "projval (ALT r1 r2) c (Right v2) = Right (projval r2 c v2)"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1077
| "projval (SEQ r1 r2) c (Seq v1 v2) = 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1078
     (if flat v1 = [] then Right(projval r2 c v2) 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1079
      else if nullable r1 then Left (Seq (projval r1 c v1) v2)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1080
                          else Seq (projval r1 c v1) v2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1081
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1082
text {*
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1083
  Injection value is related to r
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1084
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1085
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1086
lemma v3:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1087
  assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1088
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1089
apply(induct arbitrary: v rule: der.induct)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1090
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1091
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1092
apply(simp_all)[5]
49
c616ec6b1e3c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
  1093
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1094
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1095
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1096
apply(case_tac "c = c'")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1097
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1098
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1099
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1100
apply (metis Prf.intros(5))
49
c616ec6b1e3c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
  1101
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1102
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1103
apply(simp_all)[5]
49
c616ec6b1e3c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
  1104
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1105
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1106
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1107
apply (metis Prf.intros(2))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1108
apply (metis Prf.intros(3))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1109
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1110
apply(case_tac "nullable r1")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1111
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1112
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1113
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1114
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1115
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1116
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1117
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1118
apply (metis Prf.intros(1))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1119
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1120
apply (metis Prf.intros(1) mkeps_nullable)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1121
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1122
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1123
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1124
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1125
apply(rule Prf.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1126
apply(auto)[2]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1127
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1128
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1129
lemma v3s:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1130
  assumes "\<Turnstile>s v : der c r" shows "\<Turnstile>(c#s) (injval r c v) : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1131
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1132
apply(induct arbitrary: s v rule: der.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1133
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1134
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1135
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1136
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1137
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1138
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1139
apply(case_tac "c = c'")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1140
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1141
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1142
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1143
apply (metis Prfs.intros(5))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1144
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1145
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1146
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1147
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1148
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1149
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1150
apply (metis Prfs.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1151
apply (metis Prfs.intros(3))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1152
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1153
apply(case_tac "nullable r1")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1154
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1155
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1156
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1157
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1158
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1159
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1160
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1161
apply (metis Prfs.intros(1) append_Cons)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1162
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1163
apply (metis Prfs.intros(1) append_Nil mkeps_nullable_s)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1164
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1165
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1166
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1167
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1168
by (metis Prfs.intros(1) append_Cons)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1169
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1170
lemma v3n:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1171
  assumes "\<TTurnstile>n v : der c r" shows "\<TTurnstile>(Suc n) (injval r c v) : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1172
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1173
apply(induct arbitrary: n v rule: der.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1174
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1175
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1176
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1177
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1178
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1179
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1180
apply(case_tac "c = c'")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1181
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1182
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1183
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1184
apply (metis One_nat_def Prfn.intros(5))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1185
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1186
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1187
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1188
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1189
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1190
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1191
apply (metis Prfn.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1192
apply (metis Prfn.intros(3))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1193
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1194
apply(case_tac "nullable r1")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1195
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1196
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1197
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1198
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1199
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1200
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1201
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1202
apply (metis Prfn.intros(1) add.commute add_Suc_right)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1203
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1204
apply (metis Prfn.intros(1) mkeps_nullable_n plus_nat.add_0)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1205
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1206
apply(erule Prfn.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1207
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1208
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1209
by (metis Prfn.intros(1) add_Suc)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1210
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1211
lemma v3_proj:
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1212
  assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1213
  shows "\<turnstile> (projval r c v) : der c r"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1214
using assms
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1215
apply(induct rule: Prf.induct)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1216
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1217
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1218
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1219
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1220
apply (metis Prf.intros(4))
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1221
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1222
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1223
apply (metis Prf.intros(2))
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1224
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1225
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1226
apply (metis Prf.intros(3))
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1227
apply(auto)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1228
apply(rule Prf.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1229
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1230
apply (metis Prf_flat_L nullable_correctness)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1231
apply(rule Prf.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1232
apply(rule Prf.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1233
apply (metis Cons_eq_append_conv)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1234
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1235
apply(rule Prf.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1236
apply (metis Cons_eq_append_conv)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1237
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1238
done
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1239
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1240
lemma v3s_proj:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1241
  assumes "\<Turnstile>(c#s) v : r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1242
  shows "\<Turnstile>s (projval r c v) : der c r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1243
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1244
apply(induct s\<equiv>"c#s" v r arbitrary: s rule: Prfs.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1245
prefer 4
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1246
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1247
apply (metis Prfs.intros(4))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1248
prefer 2
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1249
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1250
apply (metis Prfs.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1251
prefer 2
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1252
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1253
apply (metis Prfs.intros(3))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1254
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1255
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1256
apply (metis Prfs_flat append_Nil)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1257
prefer 2
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1258
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1259
apply(subst (asm) append_eq_Cons_conv)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1260
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1261
apply (metis Prfs_flat)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1262
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1263
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1264
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1265
apply(subst (asm) append_eq_Cons_conv)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1266
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1267
apply (metis Prf_flat_L Prfs_Prf nullable_correctness)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1268
apply (metis Prfs_flat list.distinct(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1269
apply(subst (asm) append_eq_Cons_conv)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1270
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1271
apply (metis Prfs_flat)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1272
by (metis Prfs.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1273
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1274
text {*
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1275
  The string behind the injection value is an added c
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1276
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1277
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1278
lemma v4s:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1279
  assumes "\<Turnstile>s v : der c r" shows "flat (injval r c v) = c # (flat v)"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1280
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1281
apply(induct arbitrary: s v rule: der.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1282
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1283
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1284
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1285
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1286
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1287
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1288
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1289
apply(case_tac "c = c'")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1290
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1291
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1292
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1293
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1294
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1295
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1296
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1297
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1298
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1299
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1300
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1301
apply(case_tac "nullable r1")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1302
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1303
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1304
apply(simp_all (no_asm_use))[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1305
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1306
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1307
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1308
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1309
apply(simp only: injval.simps flat.simps)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1310
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1311
apply (metis mkeps_flat)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1312
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1313
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1314
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1315
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1316
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1317
lemma v4:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1318
  assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1319
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1320
apply(induct arbitrary: v rule: der.induct)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1321
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1322
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1323
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1324
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1325
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1326
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1327
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1328
apply(case_tac "c = c'")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1329
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1330
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1331
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1332
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1333
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1334
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1335
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1336
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1337
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1338
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1339
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1340
apply(case_tac "nullable r1")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1341
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1342
apply(erule Prf.cases)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1343
apply(simp_all (no_asm_use))[5]
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1344
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1345
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1346
apply(simp_all)[5]
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1347
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1348
apply(simp only: injval.simps flat.simps)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1349
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1350
apply (metis mkeps_flat)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1351
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1352
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1353
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1354
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1355
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1356
lemma v4_flats:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1357
  assumes "\<turnstile> v : der c r" "\<not>nullable r" shows "hd (flats (injval r c v)) \<noteq> []"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1358
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1359
apply(induct arbitrary: v rule: der.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1360
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1361
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1362
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1363
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1364
apply(case_tac "c = c'")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1365
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1366
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1367
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1368
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1369
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1370
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1371
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1372
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1373
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1374
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1375
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1376
apply(case_tac "nullable r1")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1377
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1378
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1379
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1380
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1381
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1382
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1383
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1384
oops
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1385
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1386
lemma v4_flat_left:
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1387
  assumes "\<turnstile> v : der c r" "\<not>(nullable_left r)" shows "flat_left (injval r c v) = c # (flat_left v)"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1388
using assms
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1389
apply(induct arbitrary: v rule: der.induct)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1390
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1391
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1392
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1393
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1394
apply(case_tac "c = c'")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1395
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1396
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1397
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1398
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1399
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1400
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1401
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1402
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1403
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1404
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1405
apply(case_tac "nullable r1")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1406
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1407
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1408
apply(simp_all (no_asm_use))[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1409
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1410
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1411
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1412
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1413
apply(simp only: injval.simps)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1414
apply (metis nullable_left)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1415
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1416
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1417
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1418
done
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1419
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1420
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1421
lemma v4_proj:
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1422
  assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1423
  shows "c # flat (projval r c v) = flat v"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1424
using assms
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1425
apply(induct rule: Prf.induct)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1426
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1427
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1428
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1429
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1430
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1431
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1432
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1433
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1434
apply(auto)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1435
by (metis Cons_eq_append_conv)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1436
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1437
lemma v4_proj2:
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1438
  assumes "\<turnstile> v : r" and "(flat v) = c # s"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1439
  shows "flat (projval r c v) = s"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1440
using assms
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1441
by (metis list.inject v4_proj)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1442
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1443
lemma injval_inj: "inj_on (injval r c) {v. \<turnstile> v : der c r}"
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1444
apply(induct c r rule: der.induct)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1445
unfolding inj_on_def
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1446
apply(auto)[1]
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1447
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1448
apply(simp_all)[5]
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1449
apply(auto)[1]
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1450
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1451
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1452
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1453
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1454
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1455
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1456
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1457
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1458
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1459
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1460
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1461
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1462
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1463
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1464
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1465
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1466
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1467
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1468
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1469
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1470
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1471
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1472
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1473
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1474
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1475
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1476
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1477
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1478
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1479
apply(clarify)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1480
apply (metis list.distinct(1) mkeps_flat v4)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1481
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1482
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1483
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1484
apply(rotate_tac 6)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1485
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1486
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1487
apply (metis list.distinct(1) mkeps_flat v4)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1488
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1489
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1490
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1491
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1492
done
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1493
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1494
lemma 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1495
  assumes "POSIXs v (der c r) s" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1496
  shows "POSIXs (injval r c v) r (c # s)"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1497
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1498
apply(induct c r arbitrary: v s rule: der.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1499
apply(auto simp add: POSIXs_def)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1500
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1501
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1502
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1503
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1504
apply(auto simp add: POSIXs_def)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1505
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1506
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1507
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1508
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1509
apply(case_tac "c = c'")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1510
apply(auto simp add: POSIXs_def)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1511
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1512
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1513
apply (metis Prfs.intros(5))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1514
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1515
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1516
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1517
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1518
apply (metis ValOrd2.intros(8))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1519
apply(auto simp add: POSIXs_def)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1520
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1521
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1522
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1523
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1524
apply(frule Prfs_POSIX)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1525
apply(drule conjunct1)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1526
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1527
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1528
apply(rule POSIXs_ALT_I1)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1529
apply (metis POSIXs_ALT2)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1530
apply(rule POSIXs_ALT_I2)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1531
apply (metis POSIXs_ALT1a)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1532
apply(frule POSIXs_ALT1b)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1533
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1534
apply(frule POSIXs_ALT1a)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1535
(* HERE *)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1536
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1537
apply(rule ccontr)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1538
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1539
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1540
apply(drule POSIXs_ALT1a)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1541
thm ValOrd2.intros
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1542
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1543
apply(simp (no_asm) add: POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1544
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1545
apply (metis POSIXs_def 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1546
der.simps(4) v3s)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1547
apply(subst (asm) (5) POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1548
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1549
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1550
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1551
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1552
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1553
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1554
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1555
apply(drule_tac x="v1a" in meta_spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1556
apply(drule_tac x="sb" in meta_spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1557
apply(drule_tac meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1558
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1559
apply (metis POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1560
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1561
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1562
apply(subst v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1563
apply (metis Prfs_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1564
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1565
apply(drule_tac x="Left (injval r1a c v1)" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1566
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1567
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1568
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1569
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1570
apply(simp_all)[8] 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1571
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1572
thm v4s
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1573
apply(subst (asm) v4s[of "sb"])
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1574
apply(assumption)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1575
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1576
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1577
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1578
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1579
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1580
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1581
apply(subst v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1582
apply (metis Prfs_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1583
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1584
apply(drule_tac x="Right (injval r2a c v2)" in spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1585
apply(drule mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1586
apply(rule Prfs.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1587
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1588
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1589
apply(simp_all)[8] 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1590
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1591
apply(subst (asm) v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1592
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1593
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1594
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1595
apply (metis POSIXs_ALT1a POSIXs_def Prfs.intros(3))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1596
apply(case_tac "nullable r1")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1597
apply(simp (no_asm) add: POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1598
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1599
apply(subst (asm) (6) POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1600
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1601
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1602
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1603
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1604
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1605
apply (metis Prfs.intros(3) der.simps(5) injval.simps(6) v3s)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1606
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1607
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1608
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1609
apply(subst (asm) (6) POSIXs_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1610
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1611
apply(rotate_tac 7)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1612
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1613
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1614
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1615
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1616
apply(rotate_tac 8)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1617
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1618
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1619
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1620
apply(rule ValOrd2.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1621
apply(drule_tac x="v1b" in meta_spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1622
apply(drule_tac x="s1a" in meta_spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1623
apply(drule meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1624
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1625
apply(subst (asm) Cons_eq_append_conv)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1626
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1627
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1628
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1629
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1630
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1631
thm v4
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1632
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1633
lemma t: "(c#xs = c#ys) \<Longrightarrow> xs = ys"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1634
by (metis list.sel(3))
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1635
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1636
lemma t2: "(xs = ys) \<Longrightarrow> (c#xs) = (c#ys)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1637
by (metis)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1638
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1639
fun zeroable where
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1640
  "zeroable NULL = True"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1641
| "zeroable EMPTY = False"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1642
| "zeroable (CHAR c) = False"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1643
| "zeroable (ALT r1 r2) = (zeroable r1 \<and> zeroable r2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1644
| "zeroable (SEQ r1 r2) = (zeroable r1 \<or> zeroable r2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1645
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1646
lemma "\<not>(nullable r) \<Longrightarrow> \<not>(\<exists>v. \<turnstile> v : r \<and> flat v = [])"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1647
by (metis Prf_flat_L nullable_correctness)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1648
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1649
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1650
lemma LeftRight:
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1651
  assumes "(Left v1) \<succ>(der c (ALT r1 r2)) (Right v2)"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1652
  and "\<turnstile> v1 : der c r1" "\<turnstile> v2 : der c r2" 
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1653
  shows "(injval (ALT r1 r2) c (Left v1)) \<succ>(ALT r1 r2) (injval (ALT r1 r2) c (Right v2))"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1654
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1655
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1656
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1657
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1658
apply(rule ValOrd.intros)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1659
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1660
apply(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1661
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1662
apply(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1663
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1664
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1665
done
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1666
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1667
lemma RightLeft:
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1668
  assumes "(Right v1) \<succ>(der c (ALT r1 r2)) (Left v2)"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1669
  and "\<turnstile> v1 : der c r2" "\<turnstile> v2 : der c r1" 
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1670
  shows "(injval (ALT r1 r2) c (Right v1)) \<succ>(ALT r1 r2) (injval (ALT r1 r2) c (Left v2))"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1671
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1672
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1673
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1674
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1675
apply(rule ValOrd.intros)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1676
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1677
apply(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1678
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1679
apply(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1680
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1681
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1682
done
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1683
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1684
lemma h: 
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1685
  assumes "nullable r1" "\<turnstile> v1 : der c r1"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1686
  shows "injval r1 c v1 \<succ>r1 mkeps r1"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1687
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1688
apply(induct r1 arbitrary: v1 rule: der.induct)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1689
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1690
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1691
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1692
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1693
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1694
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1695
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1696
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1697
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1698
apply(auto)[1]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1699
apply (metis ValOrd.intros(6))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1700
apply (metis ValOrd.intros(6))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1701
apply (metis ValOrd.intros(3) le_add2 list.size(3) mkeps_flat monoid_add_class.add.right_neutral)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1702
apply(auto)[1]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1703
apply (metis ValOrd.intros(4) length_greater_0_conv list.distinct(1) list.size(3) mkeps_flat v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1704
apply (metis ValOrd.intros(4) length_greater_0_conv list.distinct(1) list.size(3) mkeps_flat v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1705
apply (metis ValOrd.intros(5))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1706
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1707
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1708
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1709
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1710
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1711
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1712
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1713
apply (metis ValOrd.intros(2) list.distinct(1) mkeps_flat v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1714
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1715
by (metis ValOrd.intros(1))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1716
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1717
lemma LeftRightSeq:
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1718
  assumes "(Left (Seq v1 v2)) \<succ>(der c (SEQ r1 r2)) (Right v3)"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1719
  and "nullable r1" "\<turnstile> v1 : der c r1"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1720
  shows "(injval (SEQ r1 r2) c (Seq v1 v2)) \<succ>(SEQ r1 r2) (injval (SEQ r1 r2) c (Right v2))"
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1721
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1722
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1723
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1724
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1725
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1726
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1727
apply(rule ValOrd.intros(2))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1728
prefer 2
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1729
apply (metis list.distinct(1) mkeps_flat v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1730
by (metis h)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1731
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1732
lemma rr1: 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1733
  assumes "\<turnstile> v : r" "\<not>nullable r" 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1734
  shows "flat v \<noteq> []"
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1735
using assms
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1736
by (metis Prf_flat_L nullable_correctness)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1737
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1738
lemma rr2: "hd (flats v) \<noteq> [] \<Longrightarrow> flats v \<noteq> []"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1739
apply(induct v)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1740
apply(auto)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1741
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1742
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1743
lemma rr3: "flats v = [] \<Longrightarrow> flat v = []"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1744
apply(induct v)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1745
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1746
done
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1747
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1748
lemma POSIXs_der:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1749
  assumes "POSIXs v (der c r) s" "\<Turnstile>s v : der c r"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1750
  shows "POSIXs (injval r c v) r (c#s)"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1751
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1752
unfolding POSIXs_def
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1753
apply(auto)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1754
thm v3s 
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1755
apply (erule v3s)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1756
apply(drule_tac x="projval r c v'" in spec)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1757
apply(drule mp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1758
thm v3s_proj
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1759
apply(rule v3s_proj)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1760
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1761
thm v3s_proj
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1762
apply(drule v3s_proj)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1763
oops
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1764
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1765
lemma Prf_inj_test:
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1766
  assumes "v1 2\<succ> v2" "\<Turnstile>s1 v1 : der c r" "\<Turnstile>s2 v2 : der c r"  
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1767
  shows "(injval r c v1) 2\<succ>  (injval r c v2)"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1768
using assms
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1769
apply(induct c r arbitrary: s1 s2 v1 v2 rule: der.induct)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1770
(* NULL case *)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1771
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1772
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1773
(* EMPTY case *)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1774
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1775
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1776
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1777
(* CHAR case *)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1778
apply(case_tac "c = c'")
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1779
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1780
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1781
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1782
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1783
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1784
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1785
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1786
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1787
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1788
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1789
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1790
(* ALT case *)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1791
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1792
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1793
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1794
apply(auto)[2]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1795
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1796
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1797
apply(auto)[2]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1798
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1799
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1800
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1801
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1802
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1803
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1804
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1805
apply(subst v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1806
apply (metis Prfs_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1807
apply(subst v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1808
apply (metis Prfs_Prf)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1809
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1810
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1811
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1812
apply(auto)[2]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1813
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1814
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1815
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1816
apply(subst v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1817
apply (metis Prfs_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1818
apply(subst v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1819
apply (metis Prfs_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1820
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1821
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1822
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1823
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1824
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1825
(* SEQ case*)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1826
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1827
apply(case_tac "nullable r1")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1828
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1829
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1830
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1831
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1832
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1833
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1834
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1835
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1836
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1837
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1838
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1839
apply(rule ValOrd2.intros)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1840
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1841
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1842
apply(auto)[1]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1843
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1844
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1845
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1846
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1847
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1848
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1849
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1850
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1851
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1852
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1853
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1854
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1855
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1856
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1857
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1858
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1859
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1860
apply (metis ValOrd2.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1861
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1862
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1863
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1864
apply(clarify)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1865
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1866
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1867
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1868
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1869
apply(rule ValOrd2.intros)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1870
apply(rule Ord1)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1871
apply(rule h)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1872
apply(simp)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1873
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1874
apply (metis Prfs_Prf)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1875
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1876
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1877
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1878
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1879
apply(rotate_tac 6)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1880
apply(erule Prfs.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1881
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1882
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1883
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1884
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1885
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1886
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1887
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1888
apply(erule ValOrd2.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1889
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1890
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1891
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1892
using injval_inj
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1893
apply(simp add: inj_on_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1894
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1895
(* SEQ nullable case *)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1896
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1897
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1898
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1899
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1900
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1901
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1902
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1903
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1904
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1905
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1906
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1907
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1908
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1909
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1910
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1911
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1912
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1913
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1914
apply(rule ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1915
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1916
apply(rule ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1917
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1918
using injval_inj
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1919
apply(simp add: inj_on_def)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1920
apply metis
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1921
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1922
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1923
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1924
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1925
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1926
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1927
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1928
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1929
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1930
apply(rule ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1931
apply (metis h)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1932
apply (metis list.distinct(1) mkeps_flat v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1933
(* last case *)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1934
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1935
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1936
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1937
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1938
apply(rotate_tac 6)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1939
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1940
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1941
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1942
prefer 2
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1943
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1944
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1945
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1946
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1947
apply(rule ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1948
apply(metis)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1949
apply(rule ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1950
prefer 2
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1951
thm mkeps_flat v4
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1952
apply (metis list.distinct(1) mkeps_flat v4)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1953
oops
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1954
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1955
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  1956
lemma Prf_inj_test:
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1957
  assumes "v1 \<succ>(der c r) v2" "\<turnstile> v1 : der c r" "\<turnstile> v2 : der c r"  
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1958
          "flat v1 = flat v2"
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1959
  shows "(injval r c v1) \<succ>r  (injval r c v2)"
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1960
using assms
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1961
apply(induct v1 arbitrary: r v2 taking: "length o flat" rule: measure_induct_rule)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1962
apply(case_tac r)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1963
(* NULL case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1964
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1965
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1966
apply(simp_all)[5]
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1967
(* EMPTY case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1968
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1969
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1970
apply(simp_all)[5]
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1971
(* CHAR case *)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1972
apply(case_tac "c = char")
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1973
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1974
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1975
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1976
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1977
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1978
apply (metis ValOrd.intros(8))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1979
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1980
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1981
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1982
(* ALT case *)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1983
prefer 2
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1984
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1985
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1986
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1987
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1988
apply(simp_all)[5]
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1989
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1990
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1991
apply(simp_all)[8]
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1992
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1993
apply (rule ValOrd.intros(6))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1994
apply(drule_tac x="v1b" in meta_spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1995
apply(drule_tac x="rexp1" in meta_spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1996
apply(drule_tac x="v1'" in meta_spec)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1997
apply(drule_tac meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1998
apply(assumption)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1999
apply(drule_tac meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2000
apply(assumption)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2001
apply(drule_tac meta_mp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2002
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2003
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2004
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2005
apply(simp_all)[8]
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2006
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2007
apply (metis ValOrd.intros(3) monoid_add_class.add.right_neutral order_refl v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2008
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2009
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2010
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2011
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2012
apply (metis RightLeft der.simps(4) injval.simps(2) injval.simps(3))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2013
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2014
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2015
apply(simp_all)[8]
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2016
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2017
apply (metis ValOrd.intros(5))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2018
(* SEQ case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2019
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2020
apply(case_tac "nullable r1")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2021
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2022
defer
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2023
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2024
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2025
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2026
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2027
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2028
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2029
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2030
apply(simp_all)[8]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2031
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2032
apply(rule ValOrd.intros)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2033
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2034
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2035
apply(rule ValOrd.intros(2))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2036
apply(rotate_tac 2)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2037
apply(drule_tac x="v1c" in meta_spec)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2038
apply(rotate_tac 9)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2039
apply(drule_tac x="v1'" in meta_spec)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2040
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2041
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2042
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2043
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2044
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2045
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2046
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2047
apply(simp)
77
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2048
apply (metis POSIX_SEQ1)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2049
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2050
apply (metis proj_inj_id)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2051
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2052
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2053
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2054
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2055
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2056
apply(rotate_tac 6)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2057
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2058
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2059
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2060
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2061
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2062
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2063
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2064
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2065
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2066
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2067
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2068
apply(rule ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2069
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2070
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2071
apply (metis POSIX_ALT2 POSIX_SEQ1 ValOrd.intros(2) proj_inj_id)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2072
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2073
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2074
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2075
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2076
apply (metis Prf.intros(1) Prf.intros(2) ValOrd.intros(2) der.simps(5) h injval.simps(5) mkeps_flat proj_inj_id projval.simps(4) val.distinct(19))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2077
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2078
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2079
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2080
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2081
apply(rotate_tac 7)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2082
apply(erule Prf.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2083
apply(simp_all)[5]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2084
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2085
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2086
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2087
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2088
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2089
apply(drule POSIX_ALT1a)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2090
apply(rule ValOrd.intros(2))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2091
defer
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2092
apply (metis list.distinct(1) mkeps_flat v4)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2093
apply(rule ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2094
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2095
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2096
apply (metis POSIX_ALT1a)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2097
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2098
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2099
apply(erule ValOrd.cases)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2100
apply(simp_all)[8]
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2101
apply(clarify)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2102
apply(rule ValOrd.intros(1))
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2103
apply(simp)
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2104
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2105
apply(subgoal_tac "flats v1c \<noteq> []")
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2106
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2107
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2108
apply(subst v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2109
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2110
apply(subst (asm) v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2111
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2112
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2113
apply(simp add: prefix_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2114
oops
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2115
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2116
lemma Prf_inj:
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2117
  assumes "v1 \<succ>(der c r) v2" "\<turnstile> v1 : der c r" "\<turnstile> v2 : der c r" (*"flat v1 = flat v2"*)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2118
  shows "(injval r c v1) \<succ>r (injval r c v2)"
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2119
using assms
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2120
apply(induct c r arbitrary: v1 v2 rule: der.induct)
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  2121
(* NULL case *)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2122
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2123
apply(erule ValOrd.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2124
apply(simp_all)[8]
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  2125
(* EMPTY case *)
72
9128b9440e93 updated R1 and notes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
  2126
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2127
apply(erule ValOrd.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2128
apply(simp_all)[8]
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  2129
(* CHAR case *)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2130
apply(case_tac "c = c'")
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2131
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2132
apply(erule ValOrd.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2133
apply(simp_all)[8]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2134
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2135
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2136
apply(erule ValOrd.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2137
apply(simp_all)[8]
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  2138
(* ALT case *)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2139
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2140
apply(erule ValOrd.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2141
apply(simp_all)[8]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2142
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2143
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2144
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2145
apply(rotate_tac 3)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2146
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2147
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2148
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2149
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2150
apply(rotate_tac 2)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2151
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2152
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2153
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2154
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2155
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2156
apply(rotate_tac 3)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2157
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2158
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2159
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2160
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2161
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2162
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2163
apply(subst v4)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2164
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2165
apply(subst v4)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2166
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2167
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2168
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2169
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2170
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2171
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2172
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2173
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2174
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2175
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2176
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2177
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2178
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2179
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2180
(* SEQ case*)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2181
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2182
apply(case_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2183
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2184
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2185
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2186
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2187
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2188
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2189
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2190
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2191
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2192
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2193
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2194
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2195
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2196
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2197
apply(rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2198
apply metis
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2199
using injval_inj
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2200
apply(simp add: inj_on_def)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2201
apply metis
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2202
(* SEQ nullable case *)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2203
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2204
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2205
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2206
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2207
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2208
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2209
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2210
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2211
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2212
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2213
apply(simp_all)[5]
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2214
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2215
apply(erule ValOrd.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2216
apply(simp_all)[8]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2217
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2218
apply(erule ValOrd.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2219
apply(simp_all)[8]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2220
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2221
apply(rule ValOrd.intros(1))
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2222
apply(simp)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2223
apply(rule ValOrd.intros(2))
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2224
apply metis
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2225
using injval_inj
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2226
apply(simp add: inj_on_def)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2227
apply metis
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2228
apply(clarify)
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2229
apply(simp)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2230
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2231
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2232
apply(clarify)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2233
apply(erule ValOrd.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2234
apply(simp_all)[8]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2235
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2236
apply(simp)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2237
apply(rule ValOrd.intros(2))
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2238
apply (metis h)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2239
apply (metis list.distinct(1) mkeps_flat v4)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2240
(* last case *)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2241
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2242
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2243
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2244
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2245
apply(rotate_tac 6)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2246
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2247
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2248
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2249
prefer 2
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2250
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2251
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2252
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2253
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2254
apply(rule ValOrd.intros(1))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2255
apply(metis)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2256
apply(rule ValOrd.intros(2))
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2257
prefer 2
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2258
thm mkeps_flat v4
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2259
apply (metis list.distinct(1) mkeps_flat v4)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2260
oops
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2261
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2262
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2263
lemma POSIX_der:
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2264
  assumes "POSIX v (der c r)" "\<turnstile> v : der c r"
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2265
  shows "POSIX (injval r c v) r"
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2266
using assms
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2267
unfolding POSIX_def
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2268
apply(auto)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2269
thm v3
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2270
apply (erule v3)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2271
thm v4
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2272
apply(subst (asm) v4)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2273
apply(assumption)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2274
apply(drule_tac x="projval r c v'" in spec)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2275
apply(drule mp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2276
apply(rule conjI)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2277
thm v3_proj
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2278
apply(rule v3_proj)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2279
apply(simp)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2280
apply(rule_tac x="flat v" in exI)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2281
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2282
thm t
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2283
apply(rule_tac c="c" in  t)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2284
apply(simp)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2285
thm v4_proj
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2286
apply(subst v4_proj)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2287
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2288
apply(rule_tac x="flat v" in exI)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2289
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2290
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2291
oops
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2292
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2293
lemma POSIX_der:
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2294
  assumes "POSIX v (der c r)" "\<turnstile> v : der c r"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2295
  shows "POSIX (injval r c v) r"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2296
using assms
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2297
apply(induct c r arbitrary: v rule: der.induct)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2298
(* null case*)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2299
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2300
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2301
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2302
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2303
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2304
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2305
(* empty case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2306
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2307
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2308
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2309
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2310
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2311
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2312
(* char case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2313
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2314
apply(case_tac "c = c'")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2315
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2316
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2317
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2318
apply (metis Prf.intros(5))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2319
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2320
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2321
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2322
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2323
apply (metis ValOrd.intros(8))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2324
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2325
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2326
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2327
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2328
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2329
(* alt case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2330
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2331
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2332
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2333
apply(simp (no_asm) add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2334
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2335
apply (metis Prf.intros(2) v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2336
apply(rotate_tac 4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2337
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2338
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2339
apply (metis POSIX_ALT2 POSIX_def ValOrd.intros(6))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2340
apply (metis ValOrd.intros(3) order_refl)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2341
apply(simp (no_asm) add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2342
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2343
apply (metis Prf.intros(3) v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2344
apply(rotate_tac 4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2345
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2346
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2347
defer
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2348
apply (metis POSIX_ALT1a POSIX_def ValOrd.intros(5))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2349
prefer 2
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2350
apply(subst (asm) (5) POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2351
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2352
apply(rotate_tac 5)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2353
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2354
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2355
apply(rule ValOrd.intros)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2356
apply(subst (asm) v4)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2357
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2358
apply(drule_tac x="Left (projval r1a c v1)" in spec)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2359
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2360
apply(drule mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2361
apply(rule conjI)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2362
apply (metis Prf.intros(2) v3_proj)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2363
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2364
apply (metis v4_proj2)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2365
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2366
apply(simp_all)[8]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2367
apply (metis less_not_refl v4_proj2)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2368
(* seq case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2369
apply(case_tac "nullable r1")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2370
defer
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2371
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2372
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2373
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2374
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2375
apply (metis Prf.intros(1) v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2376
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2377
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2378
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2379
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2380
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2381
apply(subst (asm) (3) v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2382
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2383
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2384
apply(subgoal_tac "flat v1a \<noteq> []")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2385
prefer 2
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2386
apply (metis Prf_flat_L nullable_correctness)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2387
apply(subgoal_tac "\<exists>s. flat v1a = c # s")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2388
prefer 2
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2389
apply (metis append_eq_Cons_conv)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2390
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2391
 
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2392
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2393
apply(auto)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2394
thm v3
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2395
apply (erule v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2396
thm v4
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2397
apply(subst (asm) v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2398
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2399
apply(drule_tac x="projval r c v'" in spec)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2400
apply(drule mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2401
apply(rule conjI)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2402
thm v3_proj
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2403
apply(rule v3_proj)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2404
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2405
apply(rule_tac x="flat v" in exI)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2406
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2407
thm t
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2408
apply(rule_tac c="c" in  t)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2409
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2410
thm v4_proj
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2411
apply(subst v4_proj)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2412
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2413
apply(rule_tac x="flat v" in exI)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2414
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2415
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2416
oops
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2417
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2418
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2419
lemma POSIX_ex: "\<turnstile> v : r \<Longrightarrow> \<exists>v. POSIX v r"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2420
apply(induct r arbitrary: v)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2421
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2422
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2423
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2424
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2425
apply(rule_tac x="Void" in exI)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2426
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2427
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2428
apply (metis Prf.intros(4))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2429
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2430
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2431
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2432
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2433
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2434
apply(rule_tac x="Char c" in exI)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2435
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2436
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2437
apply (metis Prf.intros(5))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2438
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2439
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2440
apply (metis ValOrd.intros(8))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2441
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2442
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2443
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2444
apply(drule_tac x="v1" in meta_spec)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2445
apply(drule_tac x="v2" in meta_spec)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2446
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2447
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2448
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2449
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2450
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2451
apply (metis POSIX_ALT_I1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2452
apply (metis POSIX_ALT_I1 POSIX_ALT_I2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2453
apply(case_tac "nullable r1a")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2454
apply(rule_tac x="Seq (mkeps r1a) va" in exI)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2455
apply(auto simp add: POSIX_def)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2456
apply (metis Prf.intros(1) mkeps_nullable)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2457
apply(simp add: mkeps_flat)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2458
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2459
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2460
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2461
apply(case_tac "mkeps r1 = v1a")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2462
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2463
apply (rule ValOrd.intros(1))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2464
apply (metis append_Nil mkeps_flat)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2465
apply (rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2466
apply(drule mkeps_POSIX)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2467
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2468
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2469
apply metis
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2470
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2471
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2472
apply(erule disjE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2473
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2474
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2475
apply(drule_tac x="v2" in spec)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2476
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2477
lemma POSIX_ex2: "\<turnstile> v : r \<Longrightarrow> \<exists>v. POSIX v r \<and> \<turnstile> v : r"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2478
apply(induct r arbitrary: v)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2479
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2480
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2481
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2482
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2483
apply(rule_tac x="Void" in exI)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2484
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2485
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2486
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2487
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2488
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2489
apply (metis Prf.intros(4))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2490
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2491
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2492
apply(rule_tac x="Char c" in exI)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2493
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2494
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2495
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2496
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2497
apply (metis ValOrd.intros(8))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2498
apply (metis Prf.intros(5))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2499
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2500
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2501
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2502
apply(drule_tac x="v1" in meta_spec)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2503
apply(drule_tac x="v2" in meta_spec)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2504
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2505
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2506
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2507
apply(rule ccontr)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2508
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2509
apply(drule_tac x="Seq v va" in spec)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2510
apply(drule mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2511
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2512
apply (metis Prf.intros(1))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2513
oops
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2514
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2515
lemma POSIX_ALT_cases:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2516
  assumes "\<turnstile> v : (ALT r1 r2)" "POSIX v (ALT r1 r2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2517
  shows "(\<exists>v1. v = Left v1 \<and> POSIX v1 r1) \<or> (\<exists>v2. v = Right v2 \<and> POSIX v2 r2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2518
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2519
apply(erule_tac Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2520
apply(simp_all)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2521
unfolding POSIX_def
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2522
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2523
apply (metis POSIX_ALT2 POSIX_def assms(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2524
by (metis POSIX_ALT1b assms(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2525
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2526
lemma POSIX_ALT_cases2:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2527
  assumes "POSIX v (ALT r1 r2)" "\<turnstile> v : (ALT r1 r2)" 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2528
  shows "(\<exists>v1. v = Left v1 \<and> POSIX v1 r1) \<or> (\<exists>v2. v = Right v2 \<and> POSIX v2 r2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2529
using assms POSIX_ALT_cases by auto
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2530
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2531
lemma Prf_flat_empty:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2532
  assumes "\<turnstile> v : r" "flat v = []"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2533
  shows "nullable r"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2534
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2535
apply(induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2536
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2537
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2538
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2539
lemma POSIX_proj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2540
  assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2541
  shows "POSIX (projval r c v) (der c r)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2542
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2543
apply(induct r c v arbitrary: rule: projval.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2544
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2545
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2546
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2547
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2548
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2549
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2550
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2551
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2552
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2553
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2554
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2555
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2556
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2557
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2558
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2559
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2560
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2561
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2562
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2563
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2564
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2565
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2566
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2567
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2568
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2569
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2570
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2571
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2572
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2573
apply(erule_tac [!] exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2574
prefer 3
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2575
apply(frule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2576
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2577
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2578
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2579
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2580
apply(case_tac "flat v1 = []")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2581
apply(subgoal_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2582
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2583
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2584
apply(rule_tac v="v1" in Prf_flat_empty)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2585
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2586
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2587
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2588
apply(frule POSIX_SEQ2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2589
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2590
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2591
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2592
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2593
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2594
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2595
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2596
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2597
apply(rule ccontr)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2598
apply(subgoal_tac "\<turnstile> val.Right (projval r2 c v2) : (ALT (SEQ (der c r1) r2) (der c r2))")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2599
apply(rotate_tac 11)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2600
apply(frule POSIX_ex)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2601
apply(erule exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2602
apply(drule POSIX_ALT_cases2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2603
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2604
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2605
apply(drule v3_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2606
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2607
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2608
apply(drule POSIX_ex)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2609
apply(erule exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2610
apply(frule POSIX_ALT_cases2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2611
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2612
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2613
apply(erule 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2614
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2615
apply(case_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2616
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2617
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2618
apply(rotate_tac 1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2619
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2620
apply(rule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2621
apply(assumption)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2622
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2623
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2624
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2625
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2626
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2627
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2628
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2629
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2630
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2631
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2632
apply (metis Cons_eq_append_conv)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2633
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2634
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2635
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2636
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2637
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2638
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2639
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2640
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2641
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2642
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2643
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2644
lemma POSIX_proj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2645
  assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2646
  shows "POSIX (projval r c v) (der c r)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2647
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2648
apply(induct r arbitrary: c v rule: rexp.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2649
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2650
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2651
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2652
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2653
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2654
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2655
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2656
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2657
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2658
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2659
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2660
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2661
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2662
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2663
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2664
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2665
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2666
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2667
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2668
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2669
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2670
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2671
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2672
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2673
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2674
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2675
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2676
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2677
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2678
apply(erule_tac [!] exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2679
prefer 3
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2680
apply(frule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2681
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2682
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2683
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2684
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2685
apply(case_tac "flat v1 = []")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2686
apply(subgoal_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2687
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2688
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2689
apply(rule_tac v="v1" in Prf_flat_empty)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2690
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2691
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2692
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2693
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2694
lemma POSIX_proj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2695
  assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2696
  shows "POSIX (projval r c v) (der c r)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2697
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2698
apply(induct r c v arbitrary: rule: projval.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2699
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2700
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2701
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2702
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2703
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2704
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2705
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2706
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2707
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2708
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2709
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2710
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2711
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2712
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2713
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2714
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2715
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2716
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2717
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2718
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2719
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2720
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2721
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2722
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2723
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2724
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2725
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2726
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2727
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2728
apply(erule_tac [!] exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2729
prefer 3
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2730
apply(frule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2731
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2732
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2733
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2734
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2735
apply(case_tac "flat v1 = []")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2736
apply(subgoal_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2737
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2738
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2739
apply(rule_tac v="v1" in Prf_flat_empty)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2740
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2741
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2742
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2743
apply(rule ccontr)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2744
apply(drule v3_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2745
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2746
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2747
apply(drule POSIX_ex)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2748
apply(erule exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2749
apply(frule POSIX_ALT_cases2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2750
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2751
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2752
apply(erule 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2753
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2754
apply(case_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2755
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2756
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2757
apply(rotate_tac 1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2758
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2759
apply(rule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2760
apply(assumption)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2761
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2762
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2763
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2764
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2765
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2766
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2767
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2768
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2769
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2770
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2771
apply (metis Cons_eq_append_conv)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2772
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2773
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2774
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2775
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2776
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2777
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2778
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2779
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2780
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2781
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2782
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2783
(* NULL case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2784
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2785
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2786
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2787
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2788
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2789
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2790
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2791
apply(rotate_tac 4)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2792
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2793
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2794
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2795
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2796
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2797
apply(frule POSIX_ALT1a)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2798
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2799
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2800
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2801
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2802
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2803
apply(rule POSIX_ALT_I2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2804
apply(assumption)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2805
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2806
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2807
thm v4_proj2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2808
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2809
apply(subst (asm) (13) POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2810
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2811
apply(drule_tac x="projval v2" in spec)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2812
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2813
apply(drule mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2814
apply(rule conjI)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2815
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2816
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2817
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2818
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2819
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2820
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2821
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2822
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2823
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2824
apply(subst (asm) (2) POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2825
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2826
apply (metis ValOrd.intros(5))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2827
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2828
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2829
apply(rotate_tac 3)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2830
apply(drule_tac c="c" in t2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2831
apply(subst (asm) v4_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2832
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2833
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2834
thm contrapos_np contrapos_nn
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2835
apply(erule contrapos_np)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2836
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2837
apply(subst  v4_proj2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2838
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2839
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2840
apply(subgoal_tac "\<not>(length (flat v1) < length (flat (projval r2a c v2a)))")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2841
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2842
apply(erule contrapos_nn)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2843
apply (metis nat_less_le v4_proj2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2844
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2845
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2846
apply(blast)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2847
thm contrapos_nn
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2848
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2849
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2850
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2851
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2852
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2853
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2854
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2855
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2856
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2857
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2858
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2859
apply (metis POSIX_ALT2 POSIX_def flat.simps(3))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2860
apply metis
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2861
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2862
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2863
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2864
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2865
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2866
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2867
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2868
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2869
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2870
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2871
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2872
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2873
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2874
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2875
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2876
apply (metis POSIX_ALT2 POSIX_def flat.simps(3))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2877
apply metis
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2878
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2879
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2880
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2881
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2882
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2883
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2884
(* EMPTY case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2885
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2886
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2887
apply(rotate_tac 3)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2888
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2889
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2890
apply(drule_tac c="c" in t2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2891
apply(subst (asm) v4_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2892
apply(auto)[2]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2893
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2894
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2895
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2896
(* CHAR case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2897
apply(case_tac "c = c'")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2898
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2899
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2900
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2901
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2902
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2903
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2904
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2905
(* ALT case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2906
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2907
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2908
unfolding POSIX_def
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2909
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2910
thm v4
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2911
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2912
lemma Prf_inj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2913
  assumes "v1 \<succ>(der c r) v2" "\<turnstile> v1 : der c r" "\<turnstile> v2 : der c r" "flat v1 = flat v2"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2914
  shows "(injval r c v1) \<succ>r (injval r c v2)"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2915
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2916
apply(induct arbitrary: v1 v2 rule: der.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2917
(* NULL case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2918
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2919
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2920
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2921
(* EMPTY case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2922
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2923
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2924
(* CHAR case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2925
apply(case_tac "c = c'")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2926
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2927
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2928
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2929
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2930
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2931
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2932
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2933
(* ALT case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2934
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2935
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2936
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2937
apply(rule ValOrd.intros)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2938
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2939
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2940
apply(rotate_tac 3)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2941
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2942
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2943
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2944
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2945
apply(rotate_tac 2)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2946
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2947
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2948
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2949
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2950
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2951
apply(rotate_tac 3)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2952
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2953
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2954
apply(clarify)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2955
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2956
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2957
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2958
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2959
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2960
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2961
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2962
apply(simp_all)[5]
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  2963
(* SEQ case*)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2964
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2965
apply(case_tac "nullable r1")
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2966
defer
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2967
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2968
apply(erule ValOrd.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2969
apply(simp_all)[8]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2970
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2971
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2972
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2973
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2974
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2975
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2976
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2977
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2978
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2979
apply(rule ValOrd.intros(2))
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2980
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2981
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2982
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2983
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2984
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2985
defer
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2986
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2987
apply(erule ValOrd.cases)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2988
apply(simp_all del: injval.simps)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2989
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2990
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2991
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2992
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2993
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2994
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2995
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2996
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2997
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2998
apply(simp_all)[5]
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2999
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3000
apply(rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3001
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3002
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3003
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3004
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3005
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3006
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3007
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3008
txt {*
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3009
done
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  3010
(* nullable case - unfinished *)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3011
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3012
apply(erule ValOrd.cases)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3013
apply(simp_all del: injval.simps)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3014
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3015
apply(clarify)
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  3016
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3017
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3018
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3019
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3020
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3021
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3022
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3023
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3024
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3025
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3026
apply(rule ValOrd.intros(2))
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3027
oops
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3028
*}
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  3029
oops
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3030
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  3031
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3032
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3033
text {*
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3034
  Injection followed by projection is the identity.
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3035
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3036
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3037
lemma proj_inj_id:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3038
  assumes "\<turnstile> v : der c r" 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3039
  shows "projval r c (injval r c v) = v"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3040
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3041
apply(induct r arbitrary: c v rule: rexp.induct)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3042
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3043
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3044
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3045
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3046
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3047
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3048
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3049
apply(case_tac "c = char")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3050
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3051
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3052
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3053
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3054
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3055
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3056
defer
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3057
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3058
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3059
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3060
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3061
apply(case_tac "nullable rexp1")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3062
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3063
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3064
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3065
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3066
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3067
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3068
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3069
apply (metis list.distinct(1) v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3070
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3071
apply (metis mkeps_flat)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3072
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3073
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3074
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3075
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3076
apply(simp add: v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3077
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3078
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3079
lemma "L r \<noteq> {} \<Longrightarrow> \<exists>v. POSIX3 v r"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3080
apply(induct r)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3081
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3082
apply(simp add: POSIX3_def)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3083
apply(rule_tac x="Void" in exI)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3084
apply(auto)[1]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3085
apply (metis Prf.intros(4))
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3086
apply (metis POSIX3_def flat.simps(1) mkeps.simps(1) mkeps_POSIX3 nullable.simps(2) order_refl)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3087
apply(simp add: POSIX3_def)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3088
apply(rule_tac x="Char char" in exI)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3089
apply(auto)[1]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3090
apply (metis Prf.intros(5))
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3091
apply(erule Prf.cases)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3092
apply(simp_all)[5]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3093
apply (metis ValOrd.intros(8))
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3094
apply(simp add: Sequ_def)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3095
apply(auto)[1]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3096
apply(drule meta_mp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3097
apply(auto)[2]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3098
apply(drule meta_mp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3099
apply(auto)[2]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3100
apply(rule_tac x="Seq v va" in exI)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3101
apply(simp (no_asm) add: POSIX3_def)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3102
apply(auto)[1]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3103
apply (metis POSIX3_def Prf.intros(1))
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3104
apply(erule Prf.cases)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3105
apply(simp_all)[5]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3106
apply(clarify)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3107
apply(case_tac "v  \<succ>r1a v1")
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3108
apply(rule ValOrd.intros(2))
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3109
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3110
apply(case_tac "v = v1")
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3111
apply(rule ValOrd.intros(1))
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3112
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3113
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3114
apply (metis ValOrd_refl)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3115
apply(simp add: POSIX3_def)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3116
oops
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3117
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3118
lemma "\<exists>v. POSIX v r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3119
apply(induct r)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3120
apply(rule exI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3121
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3122
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)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3123
apply(rule_tac x = "Void" in exI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3124
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3125
apply (metis POSIX_def flat.simps(1) mkeps.simps(1) mkeps_POSIX nullable.simps(2))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3126
apply(rule_tac x = "Char char" in exI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3127
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3128
apply(auto) [1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3129
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3130
apply(simp_all) [5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3131
apply (metis ValOrd.intros(8))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3132
defer
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3133
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3134
apply (metis POSIX_ALT_I1)
11
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
  3135
(* maybe it is too early to instantiate this existential quantifier *)
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
  3136
(* potentially this is the wrong POSIX value *)
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3137
apply(case_tac "r1 = NULL")
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3138
apply(simp add: POSIX_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3139
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3140
apply (metis L.simps(1) L.simps(4) Prf_flat_L mkeps_flat nullable.simps(1) nullable.simps(2) nullable_correctness seq_null(2))
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3141
apply(case_tac "r1 = EMPTY")
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3142
apply(rule_tac x = "Seq Void va" in exI )
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3143
apply(simp (no_asm) add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3144
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3145
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3146
apply(simp_all)
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3147
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3148
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3149
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3150
apply(rule ValOrd.intros(2))
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3151
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3152
apply(case_tac "\<exists>c. r1 = CHAR c")
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3153
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3154
apply(rule_tac x = "Seq (Char c) va" in exI )
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3155
apply(simp (no_asm) add: POSIX_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3156
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3157
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3158
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3159
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3160
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3161
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3162
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3163
apply(rule ValOrd.intros(2))
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3164
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3165
apply(case_tac "\<exists>r1a r1b. r1 = ALT r1a r1b")
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3166
apply(auto)
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3167
oops (* not sure if this can be proved by induction *)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3168
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3169
text {* 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3170
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3171
  HERE: Crucial lemma that does not go through in the sequence case. 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3172
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3173
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3174
lemma v5:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3175
  assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3176
  shows "POSIX (injval r c v) r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3177
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3178
apply(induct arbitrary: v rule: der.induct)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3179
(* NULL case *)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3180
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3181
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3182
apply(simp_all)[5]
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3183
(* EMPTY case *)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3184
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3185
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3186
apply(simp_all)[5]
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3187
(* CHAR case *)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3188
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3189
apply(case_tac "c = c'")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3190
apply(auto simp add: POSIX_def)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3191
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3192
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3193
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3194
apply(simp_all)[5]
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3195
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3196
apply(auto)[1]
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3197
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3198
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3199
(* base cases done *)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3200
(* ALT case *)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3201
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3202
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3203
using POSIX_ALT POSIX_ALT_I1 apply blast
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3204
apply(clarify)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3205
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3206
apply(rule POSIX_ALT_I2)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3207
apply(drule POSIX_ALT1a)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3208
apply metis
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3209
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3210
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3211
apply(assumption)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3212
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3213
apply(drule POSIX_ALT1a)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3214
apply(rotate_tac 1)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3215
apply(drule_tac x="v2" in meta_spec)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3216
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3217
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3218
apply(rotate_tac 4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3219
apply(erule Prf.cases)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3220
apply(simp_all)[5]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3221
apply(rule ValOrd.intros)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3222
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3223
apply(subst (asm) v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3224
apply(assumption)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3225
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3226
thm POSIX_ALT1a POSIX_ALT1b POSIX_ALT_I2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3227
apply(subst (asm) v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3228
apply(auto simp add: POSIX_def)[1]
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3229
apply(subgoal_tac "POSIX v2 (der c r2)")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3230
prefer 2
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3231
apply(auto simp add: POSIX_def)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3232
apply (metis POSIX_ALT1a POSIX_def flat.simps(4))
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3233
apply(frule POSIX_ALT1a)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3234
apply(drule POSIX_ALT1b)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3235
apply(rule POSIX_ALT_I2)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3236
apply(rotate_tac 1)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3237
apply(drule_tac x="v2" in meta_spec)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3238
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3239
apply(subgoal_tac "\<turnstile> Right (injval r2 c v2) : (ALT r1 r2)")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3240
prefer 2
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3241
apply (metis Prf.intros(3) v3)
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3242
apply auto[1]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3243
apply(subst v4)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3244
apply(auto)[2]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3245
apply(subst (asm) (4) POSIX_def)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3246
apply(subst (asm) v4)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3247
apply(drule_tac x="v2" in meta_spec)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3248
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3249
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3250
apply(auto)[2]
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3251
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3252
thm POSIX_ALT_I2
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3253
apply(rule POSIX_ALT_I2)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3254
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3255
apply(rule ccontr)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3256
apply(auto simp add: POSIX_def)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3257
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3258
apply(rule allI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3259
apply(rule impI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3260
apply(erule conjE)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3261
thm POSIX_ALT_I2
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3262
apply(frule POSIX_ALT1a)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3263
apply(drule POSIX_ALT1b)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3264
apply(rule POSIX_ALT_I2)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3265
apply auto[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3266
apply(subst v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3267
apply(auto)[2]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3268
apply(rotate_tac 1)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3269
apply(drule_tac x="v2" in meta_spec)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3270
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3271
apply(subst (asm) (4) POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3272
apply(subst (asm) v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3273
apply(auto)[2]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3274
(* stuck in the ALT case *)