thys/Re1.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Wed, 10 Jun 2015 14:51:35 +0100
changeset 79 ca8f9645db69
parent 78 279d0bc48308
child 81 7ac7782a7318
permissions -rw-r--r--
added frisch / cardelli paper
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
79
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    50
fun zeroable where
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    51
  "zeroable NULL = True"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    52
| "zeroable EMPTY = False"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    53
| "zeroable (CHAR c) = False"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    54
| "zeroable (ALT r1 r2) = (zeroable r1 \<and> zeroable r2)"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    55
| "zeroable (SEQ r1 r2) = (zeroable r1 \<or> zeroable r2)"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    56
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    57
lemma L_ALT_cases:
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    58
  "L (ALT r1 r2) \<noteq> {} \<Longrightarrow> (L r1 \<noteq> {}) \<or> (L r1 = {} \<and> L r2 \<noteq> {})"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    59
by(auto)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
    60
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    61
fun
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    62
 nullable :: "rexp \<Rightarrow> bool"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    63
where
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    64
  "nullable (NULL) = False"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    65
| "nullable (EMPTY) = True"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    66
| "nullable (CHAR c) = False"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    67
| "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
    68
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
70
fahad
parents: 68
diff changeset
    69
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    70
lemma nullable_correctness:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    71
  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
    72
apply (induct r) 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
    73
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
    74
done
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
section {* Values *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
datatype val = 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
  Void
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
| Char char
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
| Seq val val
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
| Right val
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
| Left val
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
    85
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
section {* The string behind a value *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
fun flat :: "val \<Rightarrow> string"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
  "flat(Void) = []"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
| "flat(Char c) = [c]"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
| "flat(Left v) = flat(v)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
| "flat(Right v) = flat(v)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
| "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
    95
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    96
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
    97
where
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    98
  "flats(Void) = [[]]"
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    99
| "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
   100
| "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
   101
| "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
   102
| "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
   103
70
fahad
parents: 68
diff changeset
   104
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
   105
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
   106
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
   107
79
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   108
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
   109
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
   110
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
   111
 "\<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
   112
| "\<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
   113
| "\<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
   114
| "\<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
   115
| "\<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
   116
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
   117
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
   118
  "\<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
   119
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
   120
apply(auto)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   121
done
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   122
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
   123
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
   124
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
   125
 "\<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
   126
| "\<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
   127
| "\<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
   128
| "\<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
   129
| "\<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
   130
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   131
lemma Prfn_flat:
79
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   132
  "\<TTurnstile>n v : r \<Longrightarrow> length (flat v) = n"
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
   133
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
   134
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
   135
done
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
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
   138
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
 "\<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
   140
| "\<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
   141
| "\<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
   142
| "\<turnstile> Void : EMPTY"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
| "\<turnstile> Char c : CHAR c"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
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
   145
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
   146
  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
   147
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
   148
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
   149
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
   150
4b4c677501e7 proved 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
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
   152
  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
   153
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
   154
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
   155
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
   156
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   157
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
   158
  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
   159
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
   160
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
   161
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
   162
4b4c677501e7 proved 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
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
   164
  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
   165
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
   166
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
   167
done
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   168
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   169
fun mkeps :: "rexp \<Rightarrow> val"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
where
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   171
  "mkeps(EMPTY) = Void"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   172
| "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
   173
| "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
   174
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   175
lemma mkeps_nullable:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   176
  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
   177
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   178
apply(induct rule: nullable.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   179
apply(auto intro: Prf.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   180
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
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 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
   183
  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
   184
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
   185
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
   186
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
   187
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
   188
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
   189
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
   190
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
   191
4b4c677501e7 proved 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
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
   193
  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
   194
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
   195
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
   196
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
   197
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
   198
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
   199
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
   200
done
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   201
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   202
lemma mkeps_flat:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   203
  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
   204
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   205
apply(induct rule: nullable.induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   206
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   207
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   208
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   209
text {*
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   210
  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
   211
  value.
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   212
*}
70
fahad
parents: 68
diff changeset
   213
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
lemma Prf_flat_L:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
  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
   216
using assms
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   217
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
   218
apply(auto simp add: Sequ_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
lemma L_flat_Prf:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
  "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
   223
apply(induct r)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
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
   225
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
   226
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
   227
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
   228
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
   229
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
   230
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
79
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   234
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   235
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
section {* Ordering of values *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
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
   239
where
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   240
  "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
   241
| "\<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
   242
| "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
   243
| "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
   244
| "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
   245
| "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
   246
| "Void \<succ>EMPTY Void"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
| "(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
   248
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   249
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
   250
where
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   251
  "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
   252
| "\<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
   253
| "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
   254
| "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
   255
| "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
   256
| "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
   257
| "Void 2\<succ> Void"
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   258
| "(Char c) 2\<succ> (Char c)"
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   259
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   260
lemma Ord1:
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   261
  "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
   262
apply(induct rule: ValOrd.induct)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   263
apply(auto intro: ValOrd2.intros)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   264
done
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   265
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   266
lemma Ord2:
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   267
  "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
   268
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
   269
apply(auto intro: ValOrd.intros)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   270
done
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
   271
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   272
lemma Ord3:
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   273
  "\<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
   274
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
   275
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
   276
done
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   277
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   278
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   279
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
   280
  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
   281
  shows "v \<succ>r v"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
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
   283
apply(induct)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   284
apply(auto intro: ValOrd.intros)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
done
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   286
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
   287
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
   288
  "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
   289
  "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
   290
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
   291
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
   292
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   293
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   294
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
   295
  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
   296
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
   297
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
   298
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
   299
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
   300
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
   301
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
   302
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
   303
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
   304
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
   305
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
   306
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
   307
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
   308
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
   309
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
   310
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
   311
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
   312
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
   313
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
   314
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
   315
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
   316
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
   317
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
   318
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
   319
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
   320
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
   321
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
   322
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
   323
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
   324
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
   325
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
   326
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
   327
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
   328
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
   329
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
   330
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
   331
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
   332
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
   333
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
   334
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
   335
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
   336
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
   337
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
   338
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
   339
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
   340
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
   341
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
   342
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
   343
4b4c677501e7 proved 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
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
   345
  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
   346
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
   347
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
   348
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
   349
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
   350
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
   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(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
   356
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
   357
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
   358
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
   359
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
   360
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
   361
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
   362
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
   363
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
   364
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
   365
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
   366
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
   367
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
   368
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
   369
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
   370
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
   371
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
   372
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
   373
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
   374
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
   375
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
   376
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
   377
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
   378
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
   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(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
   382
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
   383
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
   384
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
   385
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
   386
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
   387
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
   388
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
   389
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
   390
79
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   391
definition prefix :: "string \<Rightarrow> string \<Rightarrow> bool" ("_ \<sqsubseteq> _" [100, 100] 100)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   392
where
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   393
  "s1 \<sqsubseteq> s2 \<equiv> \<exists>s3. s1 @ s3 = s2"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   394
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   395
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   396
lemma
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   397
  "L r \<noteq> {} \<longrightarrow> (\<exists>vmax \<in> {v. \<turnstile> v : r}. \<forall>v \<in> {v. \<turnstile> v : r \<and> flat v \<sqsubseteq> flat vmax}. vmax 2\<succ> v)"
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   398
apply(induct r)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   399
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   400
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   401
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   402
apply(rule_tac x="Void" in exI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   403
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   404
apply(rule conjI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   405
apply (metis Prf.intros(4))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   406
apply(rule allI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   407
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   408
apply(erule conjE)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   409
apply(erule Prf.cases)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   410
apply(simp_all)[5]
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   411
apply (metis ValOrd2.intros(7))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   412
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   413
apply(rule_tac x="Char char" in exI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   414
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   415
apply(rule conjI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   416
apply (metis Prf.intros)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   417
apply(rule allI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   418
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   419
apply(erule conjE)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   420
apply(erule Prf.cases)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   421
apply(simp_all)[5]
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   422
apply (metis ValOrd2.intros(8))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   423
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   424
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   425
apply(simp add: Sequ_def)[1]
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   426
apply(erule exE)+
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   427
apply(erule conjE)+
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   428
apply(clarify)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   429
defer
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   430
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   431
apply(drule L_ALT_cases)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   432
apply(erule disjE)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   433
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   434
apply(erule exE)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   435
apply(clarify)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   436
apply(rule_tac x="Left vmax" in exI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   437
apply(rule conjI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   438
apply (metis Prf.intros(2))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   439
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   440
apply(rule allI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   441
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   442
apply(erule conjE)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   443
apply(rotate_tac 4)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   444
apply(erule Prf.cases)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   445
apply(simp_all)[5]
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   446
apply (metis ValOrd2.intros(6))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   447
apply(clarify)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   448
apply (metis ValOrd2.intros(3) length_append ordered_cancel_comm_monoid_diff_class.le_iff_add prefix_def)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   449
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   450
apply(clarify)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   451
apply(rule_tac x="Right vmax" in exI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   452
apply(rule conjI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   453
apply (metis Prf.intros(3))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   454
apply(rule allI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   455
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   456
apply(erule conjE)+
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   457
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   458
apply(rotate_tac 4)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   459
apply(erule Prf.cases)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   460
apply(simp_all)[5]
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   461
apply (metis Prf_flat_L empty_iff)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   462
apply (metis ValOrd2.intros(5))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   463
apply(drule mp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   464
apply (metis empty_iff)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   465
apply(drule mp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   466
apply (metis empty_iff)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   467
apply(erule exE)+
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   468
apply(erule conjE)+
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   469
apply(rule_tac x="Seq vmax vmaxa" in exI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   470
apply(rule conjI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   471
apply (metis Prf.intros(1))
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   472
apply(rule allI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   473
apply(rule impI)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   474
apply(erule conjE)+
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   475
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   476
apply(rotate_tac 6)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   477
apply(erule Prf.cases)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   478
apply(simp_all)[5]
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   479
apply(auto)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   480
apply(case_tac "vmax = v1")
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   481
apply(simp)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   482
apply (simp add: ValOrd2.intros(1) prefix_def)
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   483
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   484
lemma
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   485
  "s \<in> L r \<longrightarrow> (\<exists>vmax \<in> {v. \<turnstile> v : r \<and> flat v = s}. \<forall>v \<in> {v. \<turnstile> v : r \<and> flat v = s}. vmax 2\<succ> v)"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   486
apply(induct s arbitrary: r rule: length_induct)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   487
apply(induct_tac r rule: rexp.induct)
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
   488
apply(rule impI)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   489
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   490
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
   491
apply(rule impI)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   492
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   493
apply(rule_tac x="Void" in exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   494
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
   495
apply(rule conjI)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   496
apply (metis Prf.intros(4))
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
   497
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
   498
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
   499
apply(erule conjE)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   500
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   501
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   502
apply (metis ValOrd2.intros(7))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   503
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   504
apply(rule impI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   505
apply(rule_tac x="Char char" in exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   506
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
   507
apply(rule conjI)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   508
apply (metis Prf.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   509
apply(rule allI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   510
apply(rule impI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   511
apply(erule conjE)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   512
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   513
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   514
apply (metis ValOrd2.intros(8))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   515
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   516
apply(rule impI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   517
apply(simp add: Sequ_def)[1]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   518
apply(erule exE)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   519
apply(erule conjE)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   520
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   521
defer
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   522
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   523
apply(rule conjI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   524
apply(rule impI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   525
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   526
apply(erule exE)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   527
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   528
apply(rule_tac x="Left vmax" in exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   529
apply(rule conjI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   530
apply (metis Prf.intros(2))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   531
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
   532
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
   533
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
   534
apply(erule conjE)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   535
apply(rotate_tac 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
   536
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
   537
apply(simp_all)[5]
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   538
apply (metis ValOrd2.intros(6))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   539
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   540
apply (metis ValOrd2.intros(3) order_refl)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   541
apply(rule impI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   542
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   543
apply(erule exE)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   544
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   545
apply(rule_tac x="Right vmax" in exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   546
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   547
apply(rule conjI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   548
apply (metis Prf.intros(3))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   549
apply(rule allI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   550
apply(rule impI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   551
apply(erule conjE)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   552
apply(rotate_tac 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
   553
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
   554
apply(simp_all)[5]
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   555
apply (metis ValOrd2.intros(8))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   556
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   557
apply(rule impI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   558
apply(simp add: Sequ_def)[1]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   559
apply(erule exE)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   560
apply(erule conjE)+
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
   561
apply(clarify)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   562
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   563
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   564
inductive ValOrd3 :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ 3\<succ>_ _" [100, 100, 100] 100)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   565
where
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   566
  "\<lbrakk>v2 3\<succ>r2 v2'; \<turnstile> v1 : r1\<rbrakk> \<Longrightarrow> (Seq v1 v2) 3\<succ>(SEQ r1 r2) (Seq v1 v2')" 
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   567
| "\<lbrakk>v1 3\<succ>r1 v1'; v1 \<noteq> v1'; flat v2 = flat v2'; \<turnstile> v2 : r2; \<turnstile> v2' : r2\<rbrakk> 
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   568
      \<Longrightarrow> (Seq v1 v2) 3\<succ>(SEQ r1 r2) (Seq v1' v2')" 
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   569
| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) 3\<succ>(ALT r1 r2) (Right v2)"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   570
| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) 3\<succ>(ALT r1 r2) (Left v1)"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   571
| "v2 3\<succ>r2 v2' \<Longrightarrow> (Right v2) 3\<succ>(ALT r1 r2) (Right v2')"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   572
| "v1 3\<succ>r1 v1' \<Longrightarrow> (Left v1) 3\<succ>(ALT r1 r2) (Left v1')"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   573
| "Void 3\<succ>EMPTY Void"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   574
| "(Char c) 3\<succ>(CHAR c) (Char c)"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   575
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
   576
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   577
lemma "v1 3\<succ>r v2 \<Longrightarrow> v1 \<succ>r v2 \<and> \<turnstile> v1 : r \<and> \<turnstile> v2 : r \<and> flat v1 = flat v2"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   578
apply(induct rule: ValOrd3.induct)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   579
prefer 8
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   580
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   581
apply (metis Prf.intros(5) ValOrd.intros(8))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   582
prefer 7
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   583
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   584
apply (metis Prf.intros(4) ValOrd.intros(7))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   585
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   586
apply (metis Prf.intros(1) ValOrd.intros(1))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   587
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   588
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   589
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   590
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
   591
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   592
lemma ValOrd_trans:
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   593
  assumes "v1 \<succ>r v2" "v2 \<succ>r v3" "\<turnstile> v1 : r" "\<turnstile> v2 : r" "\<turnstile> v3 : r" 
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   594
     "flat v1 = flat v2" "flat v2 = flat v3"
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
   595
  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
   596
using assms
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   597
apply(induct r arbitrary: v1 v2 v3 s1 s2 s3)
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
   598
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
   599
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
   600
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
   601
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
   602
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
   603
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
   604
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
   605
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
   606
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
   607
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
   608
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
   609
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
   610
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
   611
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
   612
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
   613
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
   614
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
   615
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
   616
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
   617
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
   618
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
   619
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
   620
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
   621
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
   622
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
   623
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
   624
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
   625
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
   626
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
   627
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
   628
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
   629
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
   630
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
   631
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
   632
apply(auto)[1]
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   633
apply metis
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   634
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   635
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   636
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   637
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   638
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   639
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   640
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   641
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   642
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   643
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   644
apply (metis Ord1 Ord3 Prf.intros(2) ValOrd2.intros(6))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   645
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   646
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   647
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   648
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   649
apply (rule ValOrd.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   650
apply(clarify)
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
   651
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
   652
4b4c677501e7 proved 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
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   654
lemma ValOrd_max:
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   655
  shows "L r \<noteq> {} \<Longrightarrow> \<exists>v. \<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<longrightarrow> v' \<succ>r v \<longrightarrow> v = v'))" 
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   656
using assms
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   657
apply(induct r)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   658
apply(auto)[3]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   659
apply(rule_tac x="Void" in exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   660
apply(rule conjI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   661
apply (metis Prf.intros(4))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   662
apply(rule allI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   663
apply(rule impI)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   664
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   665
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   666
apply(rule_tac x="Char char" in exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   667
apply(rule conjI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   668
apply (metis Prf.intros(5))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   669
apply(rule allI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   670
apply(rule impI)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   671
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   672
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   673
apply(simp add: L.simps)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   674
apply(auto simp: Sequ_def)[1]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   675
apply(drule meta_mp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   676
apply (metis empty_iff)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   677
apply(drule meta_mp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   678
apply (metis empty_iff)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   679
apply(erule exE)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   680
apply(rule_tac x="Seq v va" in exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   681
apply(rule conjI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   682
apply (metis Prf.intros(1))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   683
apply(rule allI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   684
apply(rule impI)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   685
apply(rotate_tac 4)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   686
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   687
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   688
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   689
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   690
apply metis
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   691
apply(simp only: L.simps Lattices.bounded_lattice_bot_class.sup_eq_bot_iff HOL.de_Morgan_conj)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   692
apply(erule disjE)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   693
apply(drule meta_mp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   694
apply (metis empty_iff)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   695
apply(erule exE)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   696
apply(rule exI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   697
apply(rule conjI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   698
apply(rule Prf.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   699
apply(erule conjE)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   700
apply(assumption)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   701
apply(rule allI)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   702
apply(rule impI)+
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   703
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   704
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   705
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   706
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   707
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   708
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   709
apply(drule meta_mp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   710
apply (metis Prf_flat_L empty_iff)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   711
apply(auto)[1]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   712
apply(erule ValOrd.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   713
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   714
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   715
oops
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
   716
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   717
section {* Posix definition *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   718
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   719
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
   720
where
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   721
  "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
   722
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   723
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
   724
where
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   725
  "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
   726
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   727
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
   728
unfolding POSIX_def POSIX2_def
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   729
apply(auto)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   730
apply(rule Ord1)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   731
apply(auto)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   732
apply(rule Ord3)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   733
apply(auto)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   734
done
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   735
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
   736
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
   737
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
   738
  "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
   739
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
   740
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
   741
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
   742
  "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
   743
4b4c677501e7 proved 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
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
   745
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
   746
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
   747
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
   748
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
   749
4b4c677501e7 proved 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
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
   751
  "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
   752
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
   753
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
   754
4b4c677501e7 proved 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
4b4c677501e7 proved 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
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
   757
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
   758
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
   759
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
   760
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
   761
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
   762
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
   763
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   764
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
   765
  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
   766
  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
   767
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
   768
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
   769
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
   770
oops
11
26b8a5213355 changed theory name
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   771
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
   772
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
   773
"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
   774
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
   775
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
   776
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
   777
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
   778
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
   779
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
   780
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
   781
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
   782
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
   783
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
   784
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
   785
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
   786
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
   787
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
   788
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
   789
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
   790
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
   791
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
   792
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
   793
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
   794
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
   795
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
   796
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
   797
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
   798
apply(auto)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   799
oops
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
   800
4b4c677501e7 proved 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
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
   802
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   803
lemma POSIX_SEQ1:
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   804
  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
   805
  shows "POSIX v1 r1"
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   807
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   808
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   809
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
   810
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
   811
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
   812
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
   813
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   814
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   815
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
   816
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
   817
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   818
by (metis ValOrd_refl)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   819
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
   820
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
   821
  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
   822
  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
   823
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
   824
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
   825
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
   826
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
   827
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
   828
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
   829
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
   830
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
   831
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
   832
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
   833
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
   834
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
   835
4b4c677501e7 proved 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
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
   837
  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
   838
  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
   839
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
   840
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
   841
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
   842
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
   843
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
   844
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
   845
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
   846
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
   847
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
   848
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
   849
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
   850
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
   851
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   852
lemma POSIX_SEQ2:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   853
  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
   854
  shows "POSIX v2 r2"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   855
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   856
unfolding POSIX_def
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   857
apply(auto)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   858
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
   859
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
   860
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
   861
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
   862
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   863
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   864
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
   865
apply(simp_all)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   866
done
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   867
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
   868
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
   869
  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
   870
  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
   871
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
   872
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
   873
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
   874
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
   875
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
   876
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
   877
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
   878
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
   879
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
   880
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
   881
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
   882
4b4c677501e7 proved 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
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
   884
  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
   885
  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
   886
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
   887
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
   888
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
   889
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
   890
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
   891
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
   892
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
   893
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
   894
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
   895
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
   896
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
   897
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   898
lemma POSIX_ALT2:
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   899
  assumes "POSIX (Left v1) (ALT r1 r2)"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   900
  shows "POSIX v1 r1"
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   901
using assms
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   902
unfolding POSIX_def
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   903
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   904
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   905
apply(simp_all)[5]
12
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   906
apply(drule_tac x="Left v'" in spec)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   907
apply(simp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   908
apply(drule mp)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   909
apply(rule Prf.intros)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   910
apply(auto)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   911
apply(erule ValOrd.cases)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   912
apply(simp_all)
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   913
done
232ea10bed6f Delete POSIX_ALT2
fahadausaf <fahad.ausaf@icloud.com>
parents: 11
diff changeset
   914
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
   915
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
   916
  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
   917
  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
   918
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
   919
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
   920
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
   921
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
   922
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
   923
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
   924
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
   925
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
   926
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
   927
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
   928
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
   929
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
   930
4b4c677501e7 proved 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
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
   932
  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
   933
  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
   934
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
   935
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
   936
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
   937
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
   938
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
   939
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
   940
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
   941
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
   942
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
   943
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
   944
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
   945
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
   946
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   947
lemma POSIX_ALT1a:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   948
  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
   949
  shows "POSIX v2 r2"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   950
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   951
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   952
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   953
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
   954
apply(simp_all)[5]
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   955
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
   956
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   957
apply(drule mp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   958
apply(rule Prf.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   959
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   960
apply(erule ValOrd.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   961
apply(simp_all)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   962
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   963
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
   964
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
   965
  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
   966
  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
   967
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
   968
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
   969
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
   970
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
   971
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
   972
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
   973
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
   974
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
   975
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
   976
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
   977
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
   978
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
   979
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   980
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
   981
  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
   982
  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
   983
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
   984
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
   985
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
   986
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
   987
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
   988
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
   989
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
   990
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
   991
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
   992
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
   993
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
   994
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
   995
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   996
lemma POSIX_ALT1b:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   997
  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
   998
  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
   999
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1000
apply(drule_tac POSIX_ALT1a)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1001
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1002
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1003
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1004
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
  1005
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
  1006
  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
  1007
  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
  1008
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
  1009
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
  1010
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
  1011
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
  1012
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
  1013
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1014
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
  1015
  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
  1016
  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
  1017
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
  1018
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
  1019
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
  1020
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
  1021
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
  1022
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1023
lemma POSIX_ALT_I1:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1024
  assumes "POSIX v1 r1" 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1025
  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
  1026
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1027
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1028
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1029
apply (metis Prf.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1030
apply(rotate_tac 2)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1031
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1032
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1033
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1034
apply(rule ValOrd.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1035
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1036
apply(rule ValOrd.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1037
by simp
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1038
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
  1039
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
  1040
  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
  1041
  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
  1042
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
  1043
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
  1044
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
  1045
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
  1046
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
  1047
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
  1048
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
  1049
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
  1050
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
  1051
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
  1052
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
  1053
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
  1054
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1055
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
  1056
  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
  1057
  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
  1058
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
  1059
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
  1060
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
  1061
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
  1062
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
  1063
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
  1064
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
  1065
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
  1066
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
  1067
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
  1068
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
  1069
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
  1070
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1071
lemma POSIX_ALT_I2:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1072
  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
  1073
  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
  1074
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1075
unfolding POSIX_def
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1076
apply(auto)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1077
apply (metis Prf.intros)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1078
apply(rotate_tac 3)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1079
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1080
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1081
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1082
apply(rule ValOrd.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1083
apply metis
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1084
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1085
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
  1086
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
  1087
  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
  1088
  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
  1089
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
  1090
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
  1091
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
  1092
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
  1093
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
  1094
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
  1095
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
  1096
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
  1097
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
  1098
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
  1099
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
  1100
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1101
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1102
lemma 
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1103
  "\<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
  1104
   \<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
  1105
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
  1106
apply(rule Prf.intros(3))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1107
apply(auto)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1108
apply(rotate_tac 3)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1109
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1110
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1111
apply(simp add: mkeps_flat)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1112
apply(auto)[1]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1113
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
  1114
apply(rule ValOrd.intros)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1115
apply(auto)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1116
done
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1117
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1118
lemma mkeps_POSIX:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1119
  assumes "nullable r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1120
  shows "POSIX (mkeps r) r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1121
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1122
apply(induct r)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1123
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1124
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1125
apply(auto)[1]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1126
apply (metis Prf.intros(4))
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1127
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1128
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1129
apply (metis ValOrd.intros)
66
eb97e8361211 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
  1130
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1131
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1132
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1133
apply(auto)[1]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1134
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
  1135
apply(rotate_tac 6)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1136
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1137
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1138
apply (simp add: mkeps_flat)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1139
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
  1140
apply(simp)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1141
apply (metis ValOrd.intros(1))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1142
apply (rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1143
apply metis
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1144
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
  1145
(* ALT case *)
72
9128b9440e93 updated R1 and notes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
  1146
thm mkeps.simps
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1147
apply(simp)
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1148
apply(erule disjE)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1149
apply(simp)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1150
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
  1151
(* *)
2d30c74ba67f added a section about a nullable proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
  1152
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
  1153
thm  POSIX_ALT_I1
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1154
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
  1155
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
  1156
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
  1157
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
  1158
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
  1159
apply(rotate_tac 4)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1160
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1161
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
  1162
thm mkeps_flat
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1163
apply(simp add: mkeps_flat)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1164
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
  1165
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
  1166
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
  1167
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
  1168
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
  1169
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
  1170
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
  1171
by simp
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1172
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1173
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1174
section {* Derivatives *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1175
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1176
fun
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1177
 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1178
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1179
  "der c (NULL) = NULL"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1180
| "der c (EMPTY) = NULL"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1181
| "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
  1182
| "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
  1183
| "der c (SEQ r1 r2) = 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1184
     (if nullable r1
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1185
      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
  1186
      else SEQ (der c r1) r2)"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1187
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1188
fun 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1189
 ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1190
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1191
  "ders [] r = r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1192
| "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
  1193
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1194
section {* Injection function *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1195
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1196
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
  1197
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1198
  "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
  1199
| "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
  1200
| "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
  1201
| "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
  1202
| "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
  1203
| "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
  1204
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1205
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1206
section {* Projection function *}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1207
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1208
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
  1209
where
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1210
  "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
  1211
| "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
  1212
| "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
  1213
| "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
  1214
     (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
  1215
      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
  1216
                          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
  1217
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1218
text {*
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1219
  Injection value is related to r
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1220
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1221
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1222
lemma v3:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1223
  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
  1224
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1225
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
  1226
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1227
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1228
apply(simp_all)[5]
49
c616ec6b1e3c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
  1229
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1230
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1231
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1232
apply(case_tac "c = c'")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1233
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1234
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1235
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1236
apply (metis Prf.intros(5))
49
c616ec6b1e3c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
  1237
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1238
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1239
apply(simp_all)[5]
49
c616ec6b1e3c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
  1240
apply(simp)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1241
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1242
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1243
apply (metis Prf.intros(2))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1244
apply (metis Prf.intros(3))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1245
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1246
apply(case_tac "nullable r1")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1247
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1248
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1249
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1250
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1251
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1252
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1253
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1254
apply (metis Prf.intros(1))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1255
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1256
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
  1257
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1258
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1259
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1260
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1261
apply(rule Prf.intros)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1262
apply(auto)[2]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1263
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1264
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
  1265
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
  1266
  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
  1267
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
  1268
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
  1269
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
  1270
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
  1271
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
  1272
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
  1273
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
  1274
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
  1275
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
  1276
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
  1277
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
  1278
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
  1279
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
  1280
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
  1281
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
  1282
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
  1283
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
  1284
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
  1285
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
  1286
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
  1287
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
  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 "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
  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(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
  1292
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
  1293
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
  1294
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
  1295
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
  1296
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
  1297
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
  1298
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
  1299
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
  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(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
  1302
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
  1303
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
  1304
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
  1305
4b4c677501e7 proved 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
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
  1307
  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
  1308
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
  1309
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
  1310
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
  1311
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
  1312
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
  1313
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
  1314
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
  1315
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
  1316
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
  1317
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
  1318
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
  1319
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
  1320
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
  1321
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
  1322
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
  1323
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
  1324
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
  1325
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
  1326
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
  1327
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
  1328
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
  1329
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
  1330
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
  1331
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
  1332
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
  1333
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
  1334
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
  1335
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
  1336
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
  1337
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
  1338
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
  1339
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
  1340
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
  1341
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
  1342
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
  1343
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
  1344
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
  1345
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
  1346
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1347
lemma v3_proj:
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1348
  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
  1349
  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
  1350
using assms
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1351
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
  1352
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1353
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1354
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1355
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1356
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
  1357
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1358
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1359
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
  1360
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1361
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1362
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
  1363
apply(auto)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1364
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
  1365
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1366
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
  1367
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
  1368
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
  1369
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
  1370
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1371
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
  1372
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
  1373
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1374
done
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1375
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
  1376
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
  1377
  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
  1378
  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
  1379
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
  1380
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
  1381
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
  1382
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
  1383
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
  1384
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
  1385
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
  1386
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
  1387
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
  1388
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
  1389
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
  1390
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
  1391
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
  1392
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
  1393
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
  1394
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
  1395
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
  1396
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
  1397
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
  1398
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
  1399
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
  1400
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
  1401
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
  1402
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
  1403
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
  1404
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
  1405
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
  1406
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
  1407
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
  1408
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
  1409
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1410
text {*
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1411
  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
  1412
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1413
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
  1414
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
  1415
  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
  1416
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
  1417
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
  1418
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
  1419
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
  1420
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
  1421
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
  1422
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
  1423
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
  1424
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
  1425
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
  1426
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
  1427
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
  1428
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
  1429
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
  1430
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
  1431
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
  1432
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
  1433
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
  1434
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
  1435
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
  1436
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
  1437
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
  1438
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
  1439
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
  1440
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
  1441
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
  1442
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
  1443
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
  1444
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
  1445
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
  1446
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
  1447
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
  1448
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
  1449
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
  1450
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
  1451
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
  1452
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1453
lemma v4:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1454
  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
  1455
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1456
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
  1457
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1458
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1459
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1460
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1461
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1462
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1463
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1464
apply(case_tac "c = c'")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1465
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1466
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1467
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1468
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1469
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1470
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1471
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1472
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1473
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1474
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1475
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1476
apply(case_tac "nullable r1")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1477
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1478
apply(erule Prf.cases)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1479
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
  1480
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1481
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1482
apply(simp_all)[5]
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1483
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1484
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
  1485
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1486
apply (metis mkeps_flat)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1487
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1488
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1489
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1490
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1491
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
  1492
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
  1493
  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
  1494
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
  1495
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
  1496
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
  1497
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
  1498
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
  1499
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
  1500
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
  1501
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
  1502
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
  1503
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
  1504
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
  1505
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
  1506
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
  1507
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
  1508
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
  1509
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
  1510
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
  1511
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
  1512
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
  1513
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
  1514
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
  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(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
  1517
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
  1518
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
  1519
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
  1520
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
  1521
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1522
lemma v4_flat_left:
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1523
  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
  1524
using assms
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1525
apply(induct arbitrary: v rule: der.induct)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1526
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1527
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1528
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1529
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1530
apply(case_tac "c = c'")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1531
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1532
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1533
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1534
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1535
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1536
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1537
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1538
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1539
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1540
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1541
apply(case_tac "nullable r1")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1542
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1543
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1544
apply(simp_all (no_asm_use))[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1545
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1546
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1547
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1548
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1549
apply(simp only: injval.simps)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1550
apply (metis nullable_left)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1551
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1552
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1553
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1554
done
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1555
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  1556
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1557
lemma v4_proj:
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1558
  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
  1559
  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
  1560
using assms
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1561
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
  1562
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1563
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1564
prefer 4
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1565
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1566
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1567
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1568
prefer 2
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1569
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1570
apply(auto)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1571
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
  1572
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1573
lemma v4_proj2:
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1574
  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
  1575
  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
  1576
using assms
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1577
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
  1578
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1579
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
  1580
apply(induct c r rule: der.induct)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1581
unfolding inj_on_def
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1582
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
  1583
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
  1584
apply(simp_all)[5]
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1585
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
  1586
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
  1587
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
  1588
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1589
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
  1590
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
  1591
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
  1592
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
  1593
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
  1594
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
  1595
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1596
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
  1597
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
  1598
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
  1599
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
  1600
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
  1601
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
  1602
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1603
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
  1604
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
  1605
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
  1606
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
  1607
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1608
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
  1609
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
  1610
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
  1611
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
  1612
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1613
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
  1614
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
  1615
apply(clarify)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1616
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
  1617
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1618
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1619
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1620
apply(rotate_tac 6)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1621
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1622
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1623
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
  1624
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1625
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1626
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1627
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1628
done
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1629
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
  1630
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
  1631
  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
  1632
  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
  1633
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
  1634
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
  1635
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
  1636
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
  1637
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
  1638
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
  1639
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
  1640
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
  1641
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
  1642
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
  1643
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
  1644
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
  1645
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
  1646
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
  1647
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
  1648
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
  1649
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
  1650
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
  1651
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
  1652
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
  1653
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
  1654
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
  1655
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
  1656
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
  1657
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
  1658
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
  1659
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
  1660
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
  1661
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
  1662
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
  1663
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
  1664
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
  1665
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
  1666
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
  1667
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
  1668
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
  1669
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
  1670
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
  1671
(* HERE *)
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1672
oops
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
  1673
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  1674
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
  1675
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
  1676
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1677
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
  1678
by (metis)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1679
79
ca8f9645db69 added frisch / cardelli paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
  1680
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1681
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  1682
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
  1683
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
  1684
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  1685
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1686
lemma LeftRight:
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1687
  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
  1688
  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
  1689
  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
  1690
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1691
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1692
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1693
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1694
apply(rule ValOrd.intros)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1695
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1696
apply(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1697
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1698
apply(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1699
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1700
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1701
done
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1702
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1703
lemma RightLeft:
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1704
  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
  1705
  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
  1706
  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
  1707
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1708
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1709
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1710
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1711
apply(rule ValOrd.intros)
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(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1714
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1715
apply(subst v4)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1716
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1717
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1718
done
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1719
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1720
lemma h: 
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1721
  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
  1722
  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
  1723
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1724
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
  1725
apply(simp)
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(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1728
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1729
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1730
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1731
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1732
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1733
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1734
apply(auto)[1]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1735
apply (metis ValOrd.intros(6))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1736
apply (metis ValOrd.intros(6))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1737
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
  1738
apply(auto)[1]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1739
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
  1740
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
  1741
apply (metis ValOrd.intros(5))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1742
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1743
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1744
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1745
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1746
apply(erule Prf.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1747
apply(simp_all)[5]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1748
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1749
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
  1750
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1751
by (metis ValOrd.intros(1))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1752
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1753
lemma LeftRightSeq:
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1754
  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
  1755
  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
  1756
  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
  1757
using assms
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1758
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1759
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1760
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1761
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1762
apply(simp)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1763
apply(rule ValOrd.intros(2))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1764
prefer 2
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1765
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
  1766
by (metis h)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  1767
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
  1768
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
  1769
  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
  1770
  shows "flat v \<noteq> []"
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1771
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
  1772
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
  1773
4b4c677501e7 proved 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
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
  1775
apply(induct v)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1776
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
  1777
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
  1778
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  1779
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
  1780
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
  1781
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
  1782
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
  1783
4b4c677501e7 proved 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
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
  1785
  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
  1786
  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
  1787
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
  1788
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
  1789
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
  1790
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
  1791
apply (erule v3s)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1792
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
  1793
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
  1794
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
  1795
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
  1796
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
  1797
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
  1798
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
  1799
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
  1800
4b4c677501e7 proved 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
lemma Prf_inj_test:
78
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1802
  assumes "v1 2\<succ> v2" "\<turnstile> v1 : der c r" "\<turnstile> v2 : der c r" 
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1803
  "length (flat v1) = length (flat v2)"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1804
  shows "(injval r c v1) 2\<succ>  (injval r c v2)"
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1805
using assms
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1806
apply(induct c r arbitrary: v1 v2 rule: der.induct)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1807
(* NULL case *)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1808
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1809
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1810
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1811
(* EMPTY case *)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1812
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1813
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1814
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1815
(* CHAR case *)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1816
apply(case_tac "c = c'")
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1817
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1818
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1819
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1820
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1821
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1822
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1823
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1824
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1825
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1826
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1827
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1828
(* ALT case *)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1829
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1830
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1831
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1832
apply(auto)[2]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1833
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1834
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1835
apply(auto)[2]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1836
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1837
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1838
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1839
apply(auto)[1]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1840
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1841
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1842
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1843
apply(subst v4)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1844
apply metis
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1845
apply(subst v4)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1846
apply (metis)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1847
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1848
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1849
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1850
apply(auto)[2]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1851
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1852
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1853
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1854
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1855
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1856
(* SEQ case*)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1857
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1858
apply(case_tac "nullable r1")
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1859
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1860
defer
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1861
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1862
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1863
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1864
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1865
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1866
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1867
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1868
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1869
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1870
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1871
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1872
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1873
apply(auto)[1]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1874
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1875
using injval_inj
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1876
apply(simp add: inj_on_def)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1877
apply metis
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1878
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1879
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1880
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1881
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1882
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1883
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1884
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1885
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1886
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1887
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1888
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1889
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1890
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1891
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1892
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1893
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1894
apply (metis ValOrd2.intros(1))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1895
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1896
apply metis
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1897
using injval_inj
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1898
apply(simp add: inj_on_def)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1899
apply metis
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1900
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1901
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1902
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1903
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1904
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1905
apply(rule ValOrd2.intros)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1906
apply(rule Ord1)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1907
apply(rule h)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1908
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1909
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1910
apply (metis list.distinct(1) mkeps_flat v4)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1911
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1912
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1913
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1914
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1915
apply(rotate_tac 7)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1916
apply(erule Prf.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1917
apply(simp_all)[5]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1918
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1919
defer
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1920
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1921
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1922
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1923
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1924
apply (metis ValOrd2.intros(1))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1925
apply(erule ValOrd2.cases)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1926
apply(simp_all)[8]
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1927
apply(clarify)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1928
apply(simp)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1929
apply (rule_tac ValOrd2.intros(2))
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1930
prefer 2
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1931
apply (metis list.distinct(1) mkeps_flat v4)
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1932
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1933
279d0bc48308 updated the Isabelle theories with the totality proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
  1934
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
  1935
  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
  1936
  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
  1937
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
  1938
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
  1939
(* 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
  1940
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
  1941
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
  1942
(* 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
  1943
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
  1944
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
  1945
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
  1946
(* 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
  1947
apply(case_tac "c = c'")
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1948
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
  1949
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
  1950
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
  1951
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
  1952
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
  1953
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
  1954
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
  1955
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
  1956
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
  1957
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
  1958
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
  1959
(* 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
  1960
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
  1961
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
  1962
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
  1963
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
  1964
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
  1965
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
  1966
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
  1967
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
  1968
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
  1969
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
  1970
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
  1971
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
  1972
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
  1973
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
  1974
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
  1975
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
  1976
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
  1977
apply (metis Prfs_Prf)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1978
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
  1979
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
  1980
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
  1981
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
  1982
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
  1983
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
  1984
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
  1985
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
  1986
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
  1987
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
  1988
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
  1989
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
  1990
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
  1991
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
  1992
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
  1993
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
  1994
(* 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
  1995
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
  1996
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
  1997
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
  1998
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
  1999
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
  2000
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
  2001
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
  2002
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
  2003
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
  2004
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
  2005
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
  2006
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
  2007
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
  2008
apply(rule ValOrd2.intros)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  2009
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
  2010
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
  2011
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
  2012
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
  2013
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
  2014
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
  2015
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
  2016
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
  2017
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
  2018
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
  2019
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
  2020
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
  2021
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
  2022
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
  2023
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
  2024
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
  2025
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
  2026
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
  2027
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
  2028
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
  2029
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
  2030
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
  2031
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
  2032
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
  2033
apply(clarify)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  2034
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
  2035
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
  2036
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
  2037
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
  2038
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
  2039
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
  2040
apply(rule h)
76
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  2041
apply(simp)
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  2042
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
  2043
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
  2044
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
  2045
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
  2046
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
  2047
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
  2048
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
  2049
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
  2050
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
  2051
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
  2052
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
  2053
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
  2054
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
  2055
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
  2056
4b4c677501e7 proved 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 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
  2058
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
  2059
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
  2060
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
  2061
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
  2062
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
  2063
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
  2064
(* 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
  2065
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
  2066
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
  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(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
  2069
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
  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(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
  2072
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
  2073
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
  2074
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
  2075
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
  2076
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
  2077
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
  2078
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
  2079
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
  2080
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
  2081
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
  2082
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
  2083
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
  2084
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
  2085
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
  2086
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
  2087
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
  2088
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
  2089
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
  2090
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
  2091
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
  2092
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
  2093
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
  2094
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
  2095
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
  2096
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
  2097
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
  2098
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
  2099
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
  2100
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
  2101
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
  2102
(* 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
  2103
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
  2104
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
  2105
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
  2106
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
  2107
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
  2108
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
  2109
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
  2110
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
  2111
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
  2112
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
  2113
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
  2114
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
  2115
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
  2116
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
  2117
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
  2118
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
  2119
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
  2120
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
  2121
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
  2122
oops
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  2123
39cef7b9212a added an equivalent slightly simpler POSIX definition
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  2124
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2125
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
  2126
  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
  2127
          "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
  2128
  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
  2129
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
  2130
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
  2131
apply(case_tac r)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2132
(* NULL case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2133
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
  2134
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
  2135
apply(simp_all)[5]
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2136
(* EMPTY case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2137
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
  2138
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
  2139
apply(simp_all)[5]
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2140
(* 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
  2141
apply(case_tac "c = char")
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2142
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2143
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2144
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2145
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2146
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2147
apply (metis ValOrd.intros(8))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2148
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2149
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2150
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2151
(* 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
  2152
prefer 2
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2153
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2154
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2155
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2156
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2157
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
  2158
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2159
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2160
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
  2161
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
  2162
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
  2163
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
  2164
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
  2165
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
  2166
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
  2167
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
  2168
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
  2169
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
  2170
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
  2171
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2172
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2173
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2174
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
  2175
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
  2176
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
  2177
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
  2178
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
  2179
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
  2180
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
  2181
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
  2182
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2183
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2184
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
  2185
apply(clarify)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2186
apply (metis ValOrd.intros(5))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2187
(* SEQ case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2188
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2189
apply(case_tac "nullable r1")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2190
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2191
defer
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2192
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2193
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2194
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2195
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2196
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2197
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2198
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2199
apply(simp_all)[8]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2200
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2201
apply(rule ValOrd.intros)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2202
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2203
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2204
apply(rule ValOrd.intros(2))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2205
apply(rotate_tac 2)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2206
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
  2207
apply(rotate_tac 9)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2208
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
  2209
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2210
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2211
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2212
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2213
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2214
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2215
apply(drule_tac meta_mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2216
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
  2217
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
  2218
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
  2219
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
  2220
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
  2221
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
  2222
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
  2223
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
  2224
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
  2225
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
  2226
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
  2227
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
  2228
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
  2229
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
  2230
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
  2231
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
  2232
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
  2233
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
  2234
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
  2235
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
  2236
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
  2237
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
  2238
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
  2239
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
  2240
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
  2241
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
  2242
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
  2243
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
  2244
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
  2245
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
  2246
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
  2247
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
  2248
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
  2249
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
  2250
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
  2251
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
  2252
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
  2253
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
  2254
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
  2255
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
  2256
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
  2257
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
  2258
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
  2259
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
  2260
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
  2261
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
  2262
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
  2263
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
  2264
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
  2265
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
  2266
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2267
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
  2268
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
  2269
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
  2270
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
  2271
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
  2272
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
  2273
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2274
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
  2275
4b4c677501e7 proved some basic properties (totality and trichonomity) for the orderings
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
  2276
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2277
apply(subst v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2278
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2279
apply(subst (asm) v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2280
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2281
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2282
apply(simp add: prefix_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2283
oops
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2284
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2285
lemma Prf_inj:
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2286
  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
  2287
  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
  2288
using assms
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2289
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
  2290
(* NULL case *)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2291
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2292
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
  2293
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
  2294
(* EMPTY case *)
72
9128b9440e93 updated R1 and notes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
  2295
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2296
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
  2297
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
  2298
(* CHAR case *)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2299
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
  2300
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2301
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
  2302
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
  2303
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
  2304
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2305
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
  2306
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
  2307
(* ALT case *)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2308
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2309
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
  2310
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
  2311
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
  2312
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2313
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2314
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
  2315
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
  2316
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
  2317
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2318
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2319
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
  2320
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
  2321
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
  2322
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2323
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
  2324
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2325
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
  2326
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
  2327
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
  2328
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2329
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
  2330
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2331
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2332
apply(subst v4)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2333
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2334
apply(subst v4)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2335
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2336
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2337
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
  2338
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2339
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
  2340
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
  2341
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
  2342
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
  2343
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
  2344
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  2345
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
  2346
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
  2347
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
  2348
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2349
(* SEQ case*)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2350
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2351
apply(case_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2352
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2353
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2354
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2355
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2356
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2357
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2358
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2359
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2360
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2361
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2362
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2363
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2364
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2365
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2366
apply(rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2367
apply metis
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2368
using injval_inj
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2369
apply(simp add: inj_on_def)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2370
apply metis
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2371
(* SEQ nullable case *)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2372
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2373
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2374
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2375
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2376
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2377
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2378
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2379
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2380
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2381
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2382
apply(simp_all)[5]
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2383
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2384
apply(erule ValOrd.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2385
apply(simp_all)[8]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2386
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2387
apply(erule ValOrd.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2388
apply(simp_all)[8]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2389
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2390
apply(rule ValOrd.intros(1))
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2391
apply(simp)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2392
apply(rule ValOrd.intros(2))
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2393
apply metis
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2394
using injval_inj
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2395
apply(simp add: inj_on_def)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2396
apply metis
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2397
apply(clarify)
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2398
apply(simp)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2399
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2400
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2401
apply(clarify)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2402
apply(erule ValOrd.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2403
apply(simp_all)[8]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2404
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2405
apply(simp)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2406
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
  2407
apply (metis h)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2408
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
  2409
(* last case *)
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2410
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2411
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2412
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2413
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2414
apply(rotate_tac 6)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2415
apply(erule Prf.cases)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2416
apply(simp_all)[5]
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2417
apply(clarify)
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2418
prefer 2
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2419
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2420
apply(erule ValOrd.cases)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2421
apply(simp_all)[8]
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2422
apply(clarify)
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2423
apply(rule ValOrd.intros(1))
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2424
apply(metis)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2425
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
  2426
prefer 2
73
6e035162345a solved one case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
  2427
thm mkeps_flat v4
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2428
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
  2429
oops
63
498171d2379a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 62
diff changeset
  2430
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2431
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2432
lemma POSIX_der:
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2433
  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
  2434
  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
  2435
using assms
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2436
unfolding POSIX_def
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2437
apply(auto)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2438
thm v3
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2439
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
  2440
thm v4
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2441
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
  2442
apply(assumption)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2443
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
  2444
apply(drule mp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2445
apply(rule conjI)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2446
thm v3_proj
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2447
apply(rule v3_proj)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2448
apply(simp)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2449
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
  2450
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2451
thm t
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2452
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
  2453
apply(simp)
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2454
thm v4_proj
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2455
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
  2456
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2457
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
  2458
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2459
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2460
oops
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2461
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2462
lemma POSIX_der:
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2463
  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
  2464
  shows "POSIX (injval r c v) r"
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2465
using assms
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2466
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
  2467
(* null case*)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2468
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2469
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2470
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2471
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2472
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2473
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2474
(* empty case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2475
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2476
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2477
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2478
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2479
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2480
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2481
(* char case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2482
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2483
apply(case_tac "c = c'")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2484
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2485
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2486
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2487
apply (metis Prf.intros(5))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2488
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2489
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2490
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2491
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2492
apply (metis ValOrd.intros(8))
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2493
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2494
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2495
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2496
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2497
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2498
(* alt case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2499
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2500
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2501
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2502
apply(simp (no_asm) add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2503
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2504
apply (metis Prf.intros(2) v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2505
apply(rotate_tac 4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2506
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2507
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2508
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
  2509
apply (metis ValOrd.intros(3) order_refl)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2510
apply(simp (no_asm) add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2511
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2512
apply (metis Prf.intros(3) v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2513
apply(rotate_tac 4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2514
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2515
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2516
defer
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2517
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
  2518
prefer 2
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2519
apply(subst (asm) (5) POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2520
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2521
apply(rotate_tac 5)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2522
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2523
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2524
apply(rule ValOrd.intros)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2525
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
  2526
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2527
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
  2528
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2529
apply(drule mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2530
apply(rule conjI)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2531
apply (metis Prf.intros(2) v3_proj)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2532
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2533
apply (metis v4_proj2)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2534
apply(erule ValOrd.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2535
apply(simp_all)[8]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2536
apply (metis less_not_refl v4_proj2)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2537
(* seq case *)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2538
apply(case_tac "nullable r1")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2539
defer
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2540
apply(simp add: POSIX_def)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2541
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2542
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2543
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2544
apply (metis Prf.intros(1) v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2545
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2546
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2547
apply(erule Prf.cases)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2548
apply(simp_all)[5]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2549
apply(clarify)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2550
apply(subst (asm) (3) v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2551
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2552
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2553
apply(subgoal_tac "flat v1a \<noteq> []")
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2554
prefer 2
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2555
apply (metis Prf_flat_L nullable_correctness)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2556
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
  2557
prefer 2
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2558
apply (metis append_eq_Cons_conv)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2559
apply(auto)[1]
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2560
 
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2561
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2562
apply(auto)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2563
thm v3
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2564
apply (erule v3)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2565
thm v4
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2566
apply(subst (asm) v4)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2567
apply(assumption)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2568
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
  2569
apply(drule mp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2570
apply(rule conjI)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2571
thm v3_proj
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2572
apply(rule v3_proj)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2573
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2574
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
  2575
apply(simp)
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2576
thm t
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2577
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
  2578
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2579
thm v4_proj
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2580
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
  2581
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2582
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
  2583
apply(simp)
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  2584
apply(simp)
75
f95a405c3180 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 74
diff changeset
  2585
oops
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2586
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2587
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2588
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
  2589
apply(induct r arbitrary: v)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2590
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2591
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2592
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2593
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2594
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
  2595
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2596
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2597
apply (metis Prf.intros(4))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2598
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2599
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2600
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2601
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2602
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2603
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
  2604
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2605
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2606
apply (metis Prf.intros(5))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2607
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2608
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2609
apply (metis ValOrd.intros(8))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2610
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2611
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2612
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2613
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
  2614
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
  2615
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2616
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2617
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2618
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2619
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2620
apply (metis POSIX_ALT_I1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2621
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
  2622
apply(case_tac "nullable r1a")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2623
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
  2624
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
  2625
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
  2626
apply(simp add: mkeps_flat)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2627
apply(rotate_tac 7)
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(case_tac "mkeps r1 = v1a")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2631
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2632
apply (rule ValOrd.intros(1))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2633
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
  2634
apply (rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2635
apply(drule mkeps_POSIX)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2636
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2637
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2638
apply metis
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)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2641
apply(erule disjE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2642
apply(simp)
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
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
  2645
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2646
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
  2647
apply(induct r arbitrary: v)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2648
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2649
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2650
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2651
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2652
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
  2653
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2654
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2655
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2656
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2657
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2658
apply (metis Prf.intros(4))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2659
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2660
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2661
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
  2662
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2663
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2664
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2665
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2666
apply (metis ValOrd.intros(8))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2667
apply (metis Prf.intros(5))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2668
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2669
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2670
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2671
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
  2672
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
  2673
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2674
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2675
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2676
apply(rule ccontr)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2677
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2678
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
  2679
apply(drule mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2680
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2681
apply (metis Prf.intros(1))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2682
oops
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2683
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2684
lemma POSIX_ALT_cases:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2685
  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
  2686
  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
  2687
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2688
apply(erule_tac Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2689
apply(simp_all)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2690
unfolding POSIX_def
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2691
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2692
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
  2693
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
  2694
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2695
lemma POSIX_ALT_cases2:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2696
  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
  2697
  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
  2698
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
  2699
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2700
lemma Prf_flat_empty:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2701
  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
  2702
  shows "nullable r"
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2703
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2704
apply(induct)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2705
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2706
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2707
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2708
lemma POSIX_proj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2709
  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
  2710
  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
  2711
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2712
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
  2713
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2714
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2715
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2716
defer
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(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2724
apply(simp_all)[5]
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(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2728
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2729
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2730
apply(simp_all)[5]
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(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2736
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2737
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2738
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2739
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2740
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2741
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2742
apply(erule_tac [!] exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2743
prefer 3
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2744
apply(frule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2745
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2746
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2747
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2748
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2749
apply(case_tac "flat v1 = []")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2750
apply(subgoal_tac "nullable r1")
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
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2753
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
  2754
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2755
apply(simp_all)[5]
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(frule POSIX_SEQ2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2758
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2759
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2760
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2761
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2762
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2763
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2764
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2765
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2766
apply(rule ccontr)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2767
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
  2768
apply(rotate_tac 11)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2769
apply(frule POSIX_ex)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2770
apply(erule exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2771
apply(drule POSIX_ALT_cases2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2772
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2773
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2774
apply(drule v3_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2775
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2776
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2777
apply(drule POSIX_ex)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2778
apply(erule exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2779
apply(frule POSIX_ALT_cases2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2780
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2781
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2782
apply(erule 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2783
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2784
apply(case_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2785
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2786
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2787
apply(rotate_tac 1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2788
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2789
apply(rule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2790
apply(assumption)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2791
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2792
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2793
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2794
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2795
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2796
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2797
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2798
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2799
apply(rotate_tac 7)
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 (metis Cons_eq_append_conv)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2802
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2803
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2804
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2805
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2806
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2807
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2808
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2809
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2810
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2811
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2812
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2813
lemma POSIX_proj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2814
  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
  2815
  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
  2816
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2817
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
  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
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2823
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2824
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2825
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2826
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2827
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2828
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2829
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2830
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2831
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2832
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2833
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2834
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2835
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2836
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2837
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2838
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2839
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2840
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2841
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2842
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2843
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2844
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2845
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2846
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2847
apply(erule_tac [!] exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2848
prefer 3
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2849
apply(frule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2850
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2851
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2852
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2853
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2854
apply(case_tac "flat v1 = []")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2855
apply(subgoal_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2856
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2857
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2858
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
  2859
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2860
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2861
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2862
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2863
lemma POSIX_proj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2864
  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
  2865
  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
  2866
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2867
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
  2868
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2869
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2870
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2871
defer
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2872
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2873
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2874
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2875
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2876
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2877
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2878
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2879
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2880
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2881
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2882
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2883
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2884
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2885
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2886
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2887
apply(simp_all)[5]
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(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2891
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2892
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2893
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2894
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2895
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2896
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2897
apply(erule_tac [!] exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2898
prefer 3
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2899
apply(frule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2900
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2901
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2902
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2903
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2904
apply(case_tac "flat v1 = []")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2905
apply(subgoal_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2906
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2907
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2908
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
  2909
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2910
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2911
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2912
apply(rule ccontr)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2913
apply(drule v3_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2914
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2915
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2916
apply(drule POSIX_ex)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2917
apply(erule exE)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2918
apply(frule POSIX_ALT_cases2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2919
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2920
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2921
apply(erule 
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2922
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2923
apply(case_tac "nullable r1")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2924
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2925
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2926
apply(rotate_tac 1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2927
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2928
apply(rule POSIX_SEQ1)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2929
apply(assumption)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2930
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2931
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2932
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2933
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2934
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2935
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2936
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2937
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2938
apply(rotate_tac 7)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2939
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2940
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
  2941
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2942
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2943
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2944
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2945
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2946
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2947
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2948
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2949
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2950
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2951
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2952
(* NULL case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2953
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2954
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2955
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2956
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2957
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2958
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2959
apply (metis ValOrd.intros(7))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2960
apply(rotate_tac 4)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2961
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2962
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2963
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2964
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2965
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2966
apply(frule POSIX_ALT1a)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2967
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2968
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2969
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2970
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2971
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2972
apply(rule POSIX_ALT_I2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2973
apply(assumption)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2974
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2975
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2976
thm v4_proj2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2977
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2978
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
  2979
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2980
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
  2981
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2982
apply(drule mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2983
apply(rule conjI)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2984
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2985
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2986
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2987
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2988
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2989
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2990
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2991
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2992
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2993
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
  2994
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2995
apply (metis ValOrd.intros(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(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2998
apply(rotate_tac 3)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  2999
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
  3000
apply(subst (asm) v4_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3001
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3002
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3003
thm contrapos_np contrapos_nn
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3004
apply(erule contrapos_np)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3005
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3006
apply(subst  v4_proj2)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3007
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3008
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3009
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
  3010
prefer 2
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3011
apply(erule contrapos_nn)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3012
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
  3013
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3014
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3015
apply(blast)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3016
thm contrapos_nn
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3017
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3018
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3019
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3020
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3021
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3022
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3023
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3024
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3025
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3026
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3027
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3028
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
  3029
apply metis
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3030
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3031
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3032
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3033
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3034
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3035
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3036
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3037
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3038
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3039
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3040
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3041
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3042
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3043
apply(drule meta_mp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3044
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3045
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
  3046
apply metis
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3047
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3048
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3049
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3050
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3051
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3052
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3053
(* EMPTY case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3054
apply(simp add: POSIX_def)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3055
apply(auto)[1]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3056
apply(rotate_tac 3)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3057
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3058
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3059
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
  3060
apply(subst (asm) v4_proj)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3061
apply(auto)[2]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3062
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3063
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3064
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3065
(* CHAR case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3066
apply(case_tac "c = c'")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3067
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3068
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3069
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3070
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3071
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3072
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3073
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3074
(* ALT case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3075
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3076
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3077
unfolding POSIX_def
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3078
apply(auto)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3079
thm v4
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3080
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3081
lemma Prf_inj:
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3082
  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
  3083
  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
  3084
using assms
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3085
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
  3086
(* NULL case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3087
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3088
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3089
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3090
(* EMPTY case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3091
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3092
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3093
(* CHAR case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3094
apply(case_tac "c = c'")
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3095
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3096
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3097
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3098
apply(rule ValOrd.intros)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3099
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3100
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3101
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3102
(* ALT case *)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3103
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3104
apply(erule ValOrd.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3105
apply(simp_all)[8]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3106
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
  3107
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3108
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3109
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
  3110
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
  3111
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
  3112
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3113
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3114
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
  3115
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
  3116
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
  3117
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3118
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
  3119
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3120
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
  3121
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
  3122
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3123
apply(clarify)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3124
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
  3125
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
  3126
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
  3127
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3128
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
  3129
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
  3130
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
  3131
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
  3132
(* SEQ case*)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3133
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3134
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
  3135
defer
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3136
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3137
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
  3138
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
  3139
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3140
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
  3141
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
  3142
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
  3143
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
  3144
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3145
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
  3146
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3147
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3148
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
  3149
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
  3150
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
  3151
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
  3152
apply(simp_all)[5]
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3153
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3154
defer
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3155
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3156
apply(erule ValOrd.cases)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3157
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
  3158
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3159
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3160
apply(simp)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3161
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3162
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3163
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3164
apply(simp_all)[5]
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3165
apply(clarify)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3166
apply(erule Prf.cases)
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3167
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
  3168
apply(clarify)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3169
apply(rule ValOrd.intros(2))
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3170
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3171
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3172
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3173
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3174
done
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3175
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3176
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3177
txt {*
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3178
done
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  3179
(* 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
  3180
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3181
apply(erule ValOrd.cases)
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3182
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
  3183
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3184
apply(clarify)
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  3185
apply(simp)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3186
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
  3187
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
  3188
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
  3189
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
  3190
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3191
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
  3192
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
  3193
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3194
apply(simp)
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(2))
62
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3196
oops
a6bb0152ccc2 updated some rules
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 56
diff changeset
  3197
*}
55
c33cfa1e813a added some notes (still incomplete)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
  3198
oops
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3199
74
dfa9dbb8f8e6 updated from the session today
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
  3200
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3201
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3202
text {*
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3203
  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
  3204
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3205
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3206
lemma proj_inj_id:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3207
  assumes "\<turnstile> v : der c r" 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3208
  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
  3209
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3210
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
  3211
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3212
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3213
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3214
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3215
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3216
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3217
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3218
apply(case_tac "c = char")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3219
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3220
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3221
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3222
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3223
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3224
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3225
defer
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3226
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3227
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3228
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3229
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3230
apply(case_tac "nullable rexp1")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3231
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3232
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3233
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3234
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3235
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3236
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3237
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3238
apply (metis list.distinct(1) v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3239
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3240
apply (metis mkeps_flat)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3241
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3242
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3243
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3244
apply(auto)[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3245
apply(simp add: v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3246
done
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3247
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3248
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
  3249
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
  3250
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3251
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
  3252
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
  3253
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
  3254
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
  3255
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
  3256
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
  3257
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
  3258
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
  3259
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
  3260
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
  3261
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
  3262
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
  3263
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
  3264
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
  3265
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
  3266
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
  3267
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
  3268
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
  3269
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
  3270
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
  3271
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
  3272
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
  3273
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
  3274
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
  3275
apply(clarify)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3276
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
  3277
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
  3278
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3279
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
  3280
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
  3281
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3282
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3283
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
  3284
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
  3285
oops
48
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3286
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3287
lemma "\<exists>v. POSIX v r"
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3288
apply(induct r)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3289
apply(rule exI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3290
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3291
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
  3292
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
  3293
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3294
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
  3295
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
  3296
apply(simp add: POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3297
apply(auto) [1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3298
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3299
apply(simp_all) [5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3300
apply (metis ValOrd.intros(8))
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3301
defer
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3302
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3303
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
  3304
(* 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
  3305
(* 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
  3306
apply(case_tac "r1 = NULL")
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3307
apply(simp add: POSIX_def)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3308
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3309
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
  3310
apply(case_tac "r1 = EMPTY")
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3311
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
  3312
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
  3313
apply(auto)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3314
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3315
apply(simp_all)
20
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3316
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3317
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3318
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3319
apply(rule ValOrd.intros(2))
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3320
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3321
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
  3322
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3323
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
  3324
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
  3325
apply(auto)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3326
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3327
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3328
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3329
apply(erule Prf.cases)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3330
apply(simp_all)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3331
apply(auto)[1]
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3332
apply(rule ValOrd.intros(2))
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3333
apply(rule ValOrd.intros)
c11651bbebf5 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  3334
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
  3335
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
  3336
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
  3337
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3338
text {* 
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3339
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3340
  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
  3341
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3342
*}
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3343
lemma v5:
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3344
  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
  3345
  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
  3346
using assms
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3347
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
  3348
(* NULL case *)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3349
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3350
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3351
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
  3352
(* EMPTY case *)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3353
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3354
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3355
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
  3356
(* CHAR case *)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3357
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3358
apply(case_tac "c = c'")
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3359
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
  3360
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3361
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3362
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3363
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
  3364
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
  3365
apply(auto)[1]
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3366
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3367
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3368
(* base cases done *)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3369
(* ALT case *)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3370
apply(erule Prf.cases)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3371
apply(simp_all)[5]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3372
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
  3373
apply(clarify)
53
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3374
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3375
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
  3376
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
  3377
apply metis
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3378
apply(auto)[1]
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3379
apply(subst v4)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3380
apply(assumption)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3381
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3382
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
  3383
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
  3384
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
  3385
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3386
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3387
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
  3388
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
  3389
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
  3390
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
  3391
apply(simp)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3392
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
  3393
apply(assumption)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3394
apply(clarify)
38cde0214ad5 added some lemmas, attempted others
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
  3395
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
  3396
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
  3397
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
  3398
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
  3399
prefer 2
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3400
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
  3401
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
  3402
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
  3403
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
  3404
apply(rule POSIX_ALT_I2)
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3405
apply(rotate_tac 1)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3406
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
  3407
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3408
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
  3409
prefer 2
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3410
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
  3411
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
  3412
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
  3413
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
  3414
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
  3415
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
  3416
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
  3417
apply(simp)
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3418
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3419
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
  3420
652861c9473f added a function for calculating values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  3421
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
  3422
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
  3423
10
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3424
apply(rule ccontr)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3425
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
  3426
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3427
apply(rule allI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3428
apply(rule impI)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3429
apply(erule conjE)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3430
thm POSIX_ALT_I2
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3431
apply(frule POSIX_ALT1a)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3432
apply(drule POSIX_ALT1b)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3433
apply(rule POSIX_ALT_I2)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3434
apply auto[1]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3435
apply(subst v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3436
apply(auto)[2]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3437
apply(rotate_tac 1)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3438
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
  3439
apply(simp)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3440
apply(subst (asm) (4) POSIX_def)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3441
apply(subst (asm) v4)
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3442
apply(auto)[2]
14d41b5b57b3 added test version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3443
(* stuck in the ALT case *)