thys/SpecExt.thy
author Christian Urban <christian.urban@kcl.ac.uk>
Thu, 27 Jan 2022 23:25:26 +0000
changeset 397 e1b74d618f1b
parent 359 fedc16924b76
permissions -rw-r--r--
updated Sizebound4
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
   
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory SpecExt
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
     3
  imports Main "HOL-Library.Sublist" (*"~~/src/HOL/Library/Sublist"*)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
section {* Sequential Composition of Languages *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
definition
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
where 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
text {* Two Simple Properties about Sequential Composition *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
lemma Sequ_empty_string [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  shows "A ;; {[]} = A"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  and   "{[]} ;; A = A"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
by (simp_all add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
lemma Sequ_empty [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  shows "A ;; {} = {}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  and   "{} ;; A = {}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
by (simp_all add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
lemma Sequ_assoc:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  shows "(A ;; B) ;; C = A ;; (B ;; C)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
apply blast
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
by (metis append_assoc)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
lemma Sequ_Union_in:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  shows "(A ;; (\<Union>x\<in> B. C x)) = (\<Union>x\<in> B. A ;; C x)" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
by (auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
section {* Semantic Derivative (Left Quotient) of Languages *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
definition
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  "Der c A \<equiv> {s. c # s \<in> A}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
definition
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  Ders :: "string \<Rightarrow> string set \<Rightarrow> string set"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  "Ders s A \<equiv> {s'. s @ s' \<in> A}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
lemma Der_null [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  shows "Der c {} = {}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
unfolding Der_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
by auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
lemma Der_empty [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  shows "Der c {[]} = {}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
unfolding Der_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
by auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
lemma Der_char [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  shows "Der c {[d]} = (if c = d then {[]} else {})"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
unfolding Der_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
by auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
lemma Der_union [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  shows "Der c (A \<union> B) = Der c A \<union> Der c B"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
unfolding Der_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
by auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
lemma Der_UNION [simp]: 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  shows "Der c (\<Union>x\<in>A. B x) = (\<Union>x\<in>A. Der c (B x))"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
by (auto simp add: Der_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
lemma Der_Sequ [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
unfolding Der_def Sequ_def
280
c840a99a3e05 updated
cu
parents: 279
diff changeset
    74
  by (auto simp add: Cons_eq_append_conv)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
section {* Kleene Star for Languages *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
inductive_set
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
  for A :: "string set"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  start[intro]: "[] \<in> A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
| step[intro]:  "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
(* Arden's lemma *)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
lemma Star_cases:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
  shows "A\<star> = {[]} \<union> A ;; A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
unfolding Sequ_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
by (auto) (metis Star.simps)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
lemma Star_decomp: 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  assumes "c # x \<in> A\<star>" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
by (induct x\<equiv>"c # x" rule: Star.induct) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
   (auto simp add: append_eq_Cons_conv)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
lemma Star_Der_Sequ: 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
  shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
unfolding Der_def Sequ_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
by(auto simp add: Star_decomp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
lemma Der_star [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
proof -    
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
  have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"  
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
    by (simp only: Star_cases[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  also have "... = Der c (A ;; A\<star>)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
    by (simp only: Der_union Der_empty) (simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
    by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  also have "... =  (Der c A) ;; A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
    using Star_Der_Sequ by auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
qed
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
section {* Power operation for Sets *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
fun 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  Pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _" [101, 102] 101)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
   "A \<up> 0 = {[]}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
|  "A \<up> (Suc n) = A ;; (A \<up> n)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
lemma Pow_empty [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  shows "[] \<in> A \<up> n \<longleftrightarrow> (n = 0 \<or> [] \<in> A)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
by(induct n) (auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
lemma Pow_Suc_rev:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  "A \<up> (Suc n) =  (A \<up> n) ;; A"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
apply(induct n arbitrary: A)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
apply(simp_all)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
by (metis Sequ_assoc)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
lemma Pow_decomp: 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  assumes "c # x \<in> A \<up> n" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A \<up> (n - 1)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
apply(induct n) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
apply(auto simp add: Cons_eq_append_conv Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
apply(case_tac n)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
apply(blast)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
lemma Star_Pow:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  assumes "s \<in> A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  shows "\<exists>n. s \<in> A \<up> n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
apply(induct)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
apply(auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
apply(rule_tac x="Suc n" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
lemma Pow_Star:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
  assumes "s \<in> A \<up> n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  shows "s \<in> A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
apply(induct n arbitrary: s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
apply(auto simp add: Sequ_def)
359
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   166
  done
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   167
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
lemma Der_Pow_0:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  shows "Der c (A \<up> 0) = {}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
by(simp add: Der_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
lemma Der_Pow_Suc:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
  shows "Der c (A \<up> (Suc n)) = (Der c A) ;; (A \<up> n)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
unfolding Der_def Sequ_def 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
apply(auto simp add: Cons_eq_append_conv Sequ_def dest!: Pow_decomp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
apply(case_tac n)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
apply(force simp add: Sequ_def)+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
lemma Der_Pow [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
  shows "Der c (A \<up> n) = (if n = 0 then {} else (Der c A) ;; (A \<up> (n - 1)))"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
apply(case_tac n)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
apply(simp_all del: Pow.simps add: Der_Pow_0 Der_Pow_Suc)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
lemma Der_Pow_Sequ [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
  shows "Der c (A ;; A \<up> n) = (Der c A) ;; (A \<up> n)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
by (simp only: Pow.simps[symmetric] Der_Pow) (simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
lemma Pow_Sequ_Un:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
  assumes "0 < x"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  shows "(\<Union>n \<in> {..x}. (A \<up> n)) = ({[]} \<union> (\<Union>n \<in> {..x - Suc 0}. A ;; (A \<up> n)))"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
apply(smt Pow.elims Sequ_def Suc_le_mono Suc_pred atMost_iff empty_iff insert_iff mem_Collect_eq)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
apply(rule_tac x="Suc xa" in bexI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
lemma Pow_Sequ_Un2:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  assumes "0 < x"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
  shows "(\<Union>n \<in> {x..}. (A \<up> n)) = (\<Union>n \<in> {x - Suc 0..}. A ;; (A \<up> n))"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
apply(case_tac n)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
apply fastforce
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
apply(case_tac x)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
apply(auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
apply(rule_tac x="Suc xa" in bexI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
apply(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
section {* Regular Expressions *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
datatype rexp =
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  ZERO
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
| ONE
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   220
| CH char
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
| SEQ rexp rexp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
| ALT rexp rexp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
| STAR rexp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
| UPNTIMES rexp nat
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
| NTIMES rexp nat
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
| FROMNTIMES rexp nat
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
| NMTIMES rexp nat nat
359
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   228
| NOT rexp
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
section {* Semantics of Regular Expressions *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
fun
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
  L :: "rexp \<Rightarrow> string set"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  "L (ZERO) = {}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
| "L (ONE) = {[]}"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   237
| "L (CH c) = {[c]}"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
| "L (SEQ r1 r2) = (L r1) ;; (L r2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
| "L (ALT r1 r2) = (L r1) \<union> (L r2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
| "L (STAR r) = (L r)\<star>"
280
c840a99a3e05 updated
cu
parents: 279
diff changeset
   241
| "L (UPNTIMES r n) = (\<Union>i\<in>{..n} . (L r) \<up> i)"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
| "L (NTIMES r n) = (L r) \<up> n"
280
c840a99a3e05 updated
cu
parents: 279
diff changeset
   243
| "L (FROMNTIMES r n) = (\<Union>i\<in>{n..} . (L r) \<up> i)"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
| "L (NMTIMES r n m) = (\<Union>i\<in>{n..m} . (L r) \<up> i)" 
359
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   245
| "L (NOT r) = ((UNIV:: string set)  - L r)"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
section {* Nullable, Derivatives *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
fun
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
 nullable :: "rexp \<Rightarrow> bool"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  "nullable (ZERO) = False"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
| "nullable (ONE) = True"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   254
| "nullable (CH c) = False"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
| "nullable (STAR r) = True"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
| "nullable (UPNTIMES r n) = True"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
| "nullable (NTIMES r n) = (if n = 0 then True else nullable r)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
| "nullable (FROMNTIMES r n) = (if n = 0 then True else nullable r)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
| "nullable (NMTIMES r n m) = (if m < n then False else (if n = 0 then True else nullable r))"
359
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   262
| "nullable (NOT r) = (\<not> nullable r)"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
fun
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  "der c (ZERO) = ZERO"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
| "der c (ONE) = ZERO"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   269
| "der c (CH d) = (if c = d then ONE else ZERO)"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
| "der c (SEQ r1 r2) = 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
     (if nullable r1
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
      then ALT (SEQ (der c r1) r2) (der c r2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
      else SEQ (der c r1) r2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
| "der c (STAR r) = SEQ (der c r) (STAR r)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
| "der c (UPNTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (UPNTIMES r (n - 1)))"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
| "der c (NTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (NTIMES r (n - 1)))"
294
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   278
| "der c (FROMNTIMES r n) =
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   279
     (if n = 0 
42268a284ea6 updated
cu
parents: 276
diff changeset
   280
      then SEQ (der c r) (STAR r)
42268a284ea6 updated
cu
parents: 276
diff changeset
   281
      else SEQ (der c r) (FROMNTIMES r (n - 1)))"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
| "der c (NMTIMES r n m) = 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
     (if m < n then ZERO 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
      else (if n = 0 then (if m = 0 then ZERO else 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
                           SEQ (der c r) (UPNTIMES r (m - 1))) else 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
                           SEQ (der c r) (NMTIMES r (n - 1) (m - 1))))" 
359
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   287
| "der c (NOT r) = NOT (der c r)"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
294
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   289
lemma
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   290
 "L(der c (UPNTIMES r m))  =
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   291
     L(if (m = 0) then ZERO else ALT ONE (SEQ(der c r) (UPNTIMES r (m - 1))))"
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   292
  apply(case_tac m)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   293
  apply(simp)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   294
  apply(simp del: der.simps)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   295
  apply(simp only: der.simps)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   296
  apply(simp add: Sequ_def)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   297
  apply(auto)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   298
  defer
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   299
  apply blast
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   300
  oops
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   301
359
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   302
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   303
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   304
lemma 
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   305
  assumes "der c r = ONE \<or> der c r = ZERO"
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   306
  shows "L (der c (NOT r)) \<noteq> L(if (der c r = ZERO) then ONE else 
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   307
                               if (der c r = ONE) then ZERO
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   308
                               else NOT(der c r))"
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   309
  using assms
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   310
  apply(simp)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   311
  apply(auto)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   312
  done
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   313
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   314
lemma 
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   315
  "L (der c (NOT r)) = L(if (der c r = ZERO) then ONE else 
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   316
                         if (der c r = ONE) then ZERO
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   317
                         else NOT(der c r))"
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   318
  apply(simp)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   319
  apply(auto)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   320
  oops
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   321
294
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   322
lemma pow_add:
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   323
  assumes "s1 \<in> A \<up> n" "s2 \<in> A \<up> m"
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   324
  shows "s1 @ s2 \<in> A \<up> (n + m)"
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   325
  using assms
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   326
  apply(induct n arbitrary: m s1 s2)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   327
  apply(auto simp add: Sequ_def)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   328
  by blast
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   329
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   330
lemma pow_add2:
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   331
  assumes "x \<in> A \<up> (m + n)"
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   332
  shows "x \<in> A \<up> m ;; A \<up> n"
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   333
  using assms
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   334
  apply(induct m arbitrary: n x)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   335
  apply(auto simp add: Sequ_def)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   336
  by (metis append.assoc)
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   337
  
c1de75d20aa4 added_lit
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   338
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
fun 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
 ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
  "ders [] r = r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
| "ders (c # s) r = ders s (der c r)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
lemma nullable_correctness:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
by(induct r) (auto simp add: Sequ_def) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
lemma der_correctness:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
  shows "L (der c r) = Der c (L r)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
apply(induct r) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
apply(simp add: nullable_correctness del: Der_UNION)
359
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   360
      prefer 2
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   361
      apply(simp only: der.simps)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   362
      apply(case_tac "x2 = 0")
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   363
       apply(simp)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   364
      apply(simp del: Der_Sequ L.simps)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   365
      apply(subst L.simps)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   366
  apply(subst (2) L.simps)
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   367
  thm Der_UNION
fedc16924b76 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 294
diff changeset
   368
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
apply(rule impI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
apply(subst Sequ_Union_in)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
apply(subst Der_Pow_Sequ[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
apply(subst Pow.simps[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
apply(subst Der_UNION[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
apply(subst Pow_Sequ_Un)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
apply(simp only: Der_union Der_empty)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   379
    apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
   380
(* FROMNTIMES *)    
a3134f7de065 updated
cu
parents: 275
diff changeset
   381
   apply(simp add: nullable_correctness del: Der_UNION)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   382
  apply(rule conjI)
42268a284ea6 updated
cu
parents: 276
diff changeset
   383
prefer 2    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
apply(subst Sequ_Union_in)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
apply(subst Der_Pow_Sequ[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
apply(subst Pow.simps[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
apply(case_tac x2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
prefer 2
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
apply(subst Pow_Sequ_Un2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
apply(simp)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
   392
    apply(auto simp add: Sequ_def Der_def)[1]
42268a284ea6 updated
cu
parents: 276
diff changeset
   393
   apply(auto simp add: Sequ_def split: if_splits)[1]
42268a284ea6 updated
cu
parents: 276
diff changeset
   394
  using Star_Pow apply fastforce
42268a284ea6 updated
cu
parents: 276
diff changeset
   395
  using Pow_Star apply blast
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   396
(* NMTIMES *)    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
apply(simp add: nullable_correctness del: Der_UNION)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
apply(rule impI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
apply(rule conjI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
apply(rule impI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
apply(subst Sequ_Union_in)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
apply(subst Der_Pow_Sequ[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
apply(subst Pow.simps[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
apply(subst Der_UNION[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
apply(case_tac x3a)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
apply(auto simp add: Sequ_def Der_def Cons_eq_append_conv)[1]
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
apply(rule_tac x="Suc xa" in bexI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
apply(auto simp add: Sequ_def)[2]
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   411
     apply (metis append_Cons)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   412
 apply(rule_tac x="xa - 1" in bexI)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   413
     apply(auto simp add: Sequ_def)[2]
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   414
  apply (metis One_nat_def Pow_decomp)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
apply(rule impI)+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
apply(subst Sequ_Union_in)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
apply(subst Der_Pow_Sequ[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
apply(subst Pow.simps[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
apply(subst Der_UNION[symmetric])
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
apply(case_tac x2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
apply(simp del: Pow.simps)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
apply(auto simp add: Sequ_def Der_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
apply (metis One_nat_def Suc_le_D Suc_le_mono atLeastAtMost_iff diff_Suc_1 not_le)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
by fastforce
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
lemma ders_correctness:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
  shows "L (ders s r) = Ders s (L r)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
by (induct s arbitrary: r)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
   (simp_all add: Ders_def der_correctness Der_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
section {* Values *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
datatype val = 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
  Void
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
| Char char
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
| Seq val val
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
| Right val
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
| Left val
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
| Stars "val list"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   444
| Nt val
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
section {* The string behind a value *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
fun 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
  flat :: "val \<Rightarrow> string"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
  "flat (Void) = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
| "flat (Char c) = [c]"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
| "flat (Left v) = flat v"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
| "flat (Right v) = flat v"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
| "flat (Seq v1 v2) = (flat v1) @ (flat v2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
| "flat (Stars []) = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
| "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" 
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   458
| "flat (Nt v) = flat v"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
abbreviation
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
  "flats vs \<equiv> concat (map flat vs)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
lemma flat_Stars [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
 "flat (Stars vs) = flats vs"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
by (induct vs) (auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
lemma Star_concat:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
  assumes "\<forall>s \<in> set ss. s \<in> A"  
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
  shows "concat ss \<in> A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
using assms by (induct ss) (auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
lemma Star_cstring:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
  assumes "s \<in> A\<star>"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
  shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
apply(induct rule: Star.induct)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
apply(auto)[1]
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
apply(rule_tac x="[]" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
apply(erule exE)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
apply(case_tac "s1 = []")
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
apply(rule_tac x="ss" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
apply(rule_tac x="s1#ss" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
lemma Aux:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
  assumes "\<forall>s\<in>set ss. s = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
  shows "concat ss = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
by (induct ss) (auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
lemma Pow_cstring_nonempty:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
  assumes "s \<in> A \<up> n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
  shows "\<exists>ss. concat ss = s \<and> length ss \<le> n \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
apply(induct n arbitrary: s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
apply(auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
apply(simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
apply(erule exE)+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
apply(drule_tac x="s2" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
apply(case_tac "s1 = []")
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
apply(rule_tac x="ss" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
apply(rule_tac x="s1 # ss" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
lemma Pow_cstring:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
  assumes "s \<in> A \<up> n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
  shows "\<exists>ss1 ss2. concat (ss1 @ ss2) = s \<and> length (ss1 @ ss2) = n \<and> 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
         (\<forall>s \<in> set ss1. s \<in> A \<and> s \<noteq> []) \<and> (\<forall>s \<in> set ss2. s \<in> A \<and> s = [])"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
apply(induct n arbitrary: s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
apply(auto)[1]
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
apply(simp only: Pow_Suc_rev)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
apply(simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
apply(erule exE)+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
apply(drule_tac x="s1" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
apply(erule exE)+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
apply(case_tac "s2 = []")
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
apply(rule_tac x="ss1" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
apply(rule_tac x="s2#ss2" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
apply(rule_tac x="ss1 @ [s2]" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
apply(rule_tac x="ss2" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
apply(subst Aux)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
apply(auto)[1]
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
apply(subst Aux)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
apply(auto)[1]
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
section {* Lexical Values *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
inductive 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
  Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile> _ : _" [100, 100] 100)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
 "\<lbrakk>\<Turnstile> v1 : r1; \<Turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<Turnstile>  Seq v1 v2 : SEQ r1 r2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
| "\<Turnstile> v1 : r1 \<Longrightarrow> \<Turnstile> Left v1 : ALT r1 r2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
| "\<Turnstile> v2 : r2 \<Longrightarrow> \<Turnstile> Right v2 : ALT r1 r2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
| "\<Turnstile> Void : ONE"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   557
| "\<Turnstile> Char c : CH c"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : STAR r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []; length vs \<le> n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : UPNTIMES r n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
    \<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
    length (vs1 @ vs2) = n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : NTIMES r n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r  \<and> flat v \<noteq> []; 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
    \<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; 
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   565
    length (vs1 @ vs2) = n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : FROMNTIMES r n"
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   566
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r  \<and> flat v \<noteq> []; length vs > n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : FROMNTIMES r n"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> [];
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
    \<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; 
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   569
    length (vs1 @ vs2) = n; length (vs1 @ vs2) \<le> m\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : NMTIMES r n m"
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   570
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> [];
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   571
    length vs > n; length vs \<le> m\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : NMTIMES r n m"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
293
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   573
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   574
 
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   575
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   576
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
inductive_cases Prf_elims:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
  "\<Turnstile> v : ZERO"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
  "\<Turnstile> v : SEQ r1 r2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
  "\<Turnstile> v : ALT r1 r2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
  "\<Turnstile> v : ONE"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   582
  "\<Turnstile> v : CH c"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
  "\<Turnstile> vs : STAR r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
  "\<Turnstile> vs : UPNTIMES r n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
  "\<Turnstile> vs : NTIMES r n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
  "\<Turnstile> vs : FROMNTIMES r n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
  "\<Turnstile> vs : NMTIMES r n m"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
lemma Prf_Stars_appendE:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
  assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
  shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
by (auto intro: Prf.intros elim!: Prf_elims)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
274
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   595
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   596
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
lemma flats_empty:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
  assumes "(\<forall>v\<in>set vs. flat v = [])"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
  shows "flats vs = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
by(induct vs) (simp_all)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
lemma Star_cval:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
  shows "\<exists>vs. flats vs = concat ss \<and> (\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
apply(induct ss)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
apply(auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
apply(rule_tac x="[]" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
apply(case_tac "flat v = []")
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
apply(rule_tac x="vs" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
apply(rule_tac x="v#vs" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
lemma flats_cval:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
  assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
  shows "\<exists>vs1 vs2. flats (vs1 @ vs2) = concat ss \<and> length (vs1 @ vs2) = length ss \<and> 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
          (\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []) \<and>
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
          (\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = [])"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
apply(induct ss rule: rev_induct)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
apply(rule_tac x="[]" in exI)+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   628
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
apply(case_tac "flat v = []")
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
apply(rule_tac x="vs1" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   632
apply(rule_tac x="v#vs2" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
apply(rule_tac x="vs1 @ [v]" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   635
apply(rule_tac x="vs2" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
apply(subst (asm) (2) flats_empty)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
lemma flats_cval_nonempty:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
  assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
  shows "\<exists>vs. flats vs = concat ss \<and> length vs \<le> length ss \<and> 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
          (\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   646
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   647
apply(induct ss)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
apply(rule_tac x="[]" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
apply(case_tac "flat v = []")
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   653
apply(rule_tac x="vs" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   654
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   655
apply(rule_tac x="v # vs" in exI)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
lemma Pow_flats:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
  assumes "\<forall>v \<in> set vs. flat v \<in> A"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
  shows "flats vs \<in> A \<up> length vs"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
by(induct vs)(auto simp add: Sequ_def)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
lemma Pow_flats_appends:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   666
  assumes "\<forall>v \<in> set vs1. flat v \<in> A" "\<forall>v \<in> set vs2. flat v \<in> A"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
  shows "flats vs1 @ flats vs2 \<in> A \<up> (length vs1 + length vs2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   668
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
apply(induct vs1)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   670
apply(auto simp add: Sequ_def Pow_flats)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
lemma L_flat_Prf1:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
  assumes "\<Turnstile> v : r" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
  shows "flat v \<in> L r"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
apply(induct) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
apply(auto simp add: Sequ_def Star_concat Pow_flats)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
apply(meson Pow_flats atMost_iff)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
using Pow_flats_appends apply blast
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   681
using Pow_flats_appends apply blast
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   682
apply (meson Pow_flats atLeast_iff less_imp_le)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   683
apply(rule_tac x="length vs1 + length vs2" in  bexI)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
apply(meson Pow_flats_appends atLeastAtMost_iff)
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   685
apply(simp)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   686
apply(meson Pow_flats atLeastAtMost_iff less_or_eq_imp_le)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   688
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   689
lemma L_flat_Prf2:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   690
  assumes "s \<in> L r" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
  shows "\<exists>v. \<Turnstile> v : r \<and> flat v = s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
proof(induct r arbitrary: s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
  case (STAR r s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
  have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
  have "s \<in> L (STAR r)" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
  then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   698
  using Star_cstring by auto  
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
  then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
  using IH Star_cval by metis 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
  then show "\<exists>v. \<Turnstile> v : STAR r \<and> flat v = s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
  using Prf.intros(6) flat_Stars by blast
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
  case (SEQ r1 r2 s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
  then show "\<exists>v. \<Turnstile> v : SEQ r1 r2 \<and> flat v = s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
  unfolding Sequ_def L.simps by (fastforce intro: Prf.intros)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
next
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   708
  case (ALT r1 r2 s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
  then show "\<exists>v. \<Turnstile> v : ALT r1 r2 \<and> flat v = s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
  unfolding L.simps by (fastforce intro: Prf.intros)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
next
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
  case (NTIMES r n)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
  have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
  have "s \<in> L (NTIMES r n)" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
  then obtain ss1 ss2 where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = n" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
    "\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   717
  using Pow_cstring by force
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   718
  then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = n" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   719
      "\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   720
  using IH flats_cval 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   721
  apply -
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
  apply(drule_tac x="ss1 @ ss2" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
  apply(drule_tac x="r" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   724
  apply(drule meta_mp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
  apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
  apply (metis Un_iff)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   727
  apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   728
  apply(drule_tac x="vs1" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   729
  apply(drule_tac x="vs2" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   730
  apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   731
  done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
  then show "\<exists>v. \<Turnstile> v : NTIMES r n \<and> flat v = s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   733
  using Prf.intros(8) flat_Stars by blast
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   735
  case (FROMNTIMES r n)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   736
  have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
  have "s \<in> L (FROMNTIMES r n)" by fact 
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   738
  then obtain ss1 ss2 k where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = k"  "n \<le> k"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
    "\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []"
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   740
    using Pow_cstring by force 
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   741
  then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = k" "n \<le> k"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
      "\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
  using IH flats_cval 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
  apply -
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
  apply(drule_tac x="ss1 @ ss2" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
  apply(drule_tac x="r" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   747
  apply(drule meta_mp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   748
  apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
  apply (metis Un_iff)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   750
  apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
  apply(drule_tac x="vs1" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
  apply(drule_tac x="vs2" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
  apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
  done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
  then show "\<exists>v. \<Turnstile> v : FROMNTIMES r n \<and> flat v = s"
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   756
  apply(case_tac "length vs1 \<le> n")
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   757
  apply(rule_tac x="Stars (vs1 @ take (n - length vs1) vs2)" in exI)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   758
  apply(simp)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   759
  apply(subgoal_tac "flats (take (n - length vs1) vs2) = []")
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   760
  prefer 2
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   761
  apply (meson flats_empty in_set_takeD)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   762
  apply(clarify)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   763
    apply(rule conjI)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   764
      apply(rule Prf.intros)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   765
        apply(simp)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   766
       apply (meson in_set_takeD)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   767
      apply(simp)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   768
     apply(simp)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   769
     apply (simp add: flats_empty)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   770
      apply(rule_tac x="Stars vs1" in exI)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   771
  apply(simp)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   772
    apply(rule conjI)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   773
     apply(rule Prf.intros(10))
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   774
      apply(auto)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   775
  done    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   776
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   777
  case (NMTIMES r n m)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
  have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
  have "s \<in> L (NMTIMES r n m)" by fact 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   780
  then obtain ss1 ss2 k where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = k" "n \<le> k" "k \<le> m" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   781
    "\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
  using Pow_cstring by (auto, blast)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
  then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = k" "n \<le> k" "k \<le> m"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
      "\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
  using IH flats_cval 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   786
  apply -
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
  apply(drule_tac x="ss1 @ ss2" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   788
  apply(drule_tac x="r" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
  apply(drule meta_mp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   790
  apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   791
  apply (metis Un_iff)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   792
  apply(clarify)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   793
  apply(drule_tac x="vs1" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   794
  apply(drule_tac x="vs2" in meta_spec)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   795
  apply(simp)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   796
  done
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   797
  then show "\<exists>v. \<Turnstile> v : NMTIMES r n m \<and> flat v = s"
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   798
    apply(case_tac "length vs1 \<le> n")
a3134f7de065 updated
cu
parents: 275
diff changeset
   799
  apply(rule_tac x="Stars (vs1 @ take (n - length vs1) vs2)" in exI)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   800
  apply(simp)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   801
  apply(subgoal_tac "flats (take (n - length vs1) vs2) = []")
a3134f7de065 updated
cu
parents: 275
diff changeset
   802
  prefer 2
a3134f7de065 updated
cu
parents: 275
diff changeset
   803
  apply (meson flats_empty in_set_takeD)
a3134f7de065 updated
cu
parents: 275
diff changeset
   804
  apply(clarify)
a3134f7de065 updated
cu
parents: 275
diff changeset
   805
    apply(rule conjI)
a3134f7de065 updated
cu
parents: 275
diff changeset
   806
      apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
   807
        apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
   808
       apply (meson in_set_takeD)
a3134f7de065 updated
cu
parents: 275
diff changeset
   809
      apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
   810
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
   811
     apply (simp add: flats_empty)
a3134f7de065 updated
cu
parents: 275
diff changeset
   812
      apply(rule_tac x="Stars vs1" in exI)
a3134f7de065 updated
cu
parents: 275
diff changeset
   813
  apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
   814
    apply(rule conjI)
a3134f7de065 updated
cu
parents: 275
diff changeset
   815
     apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
   816
      apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
   817
  done    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   818
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
  case (UPNTIMES r n s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
  have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
  have "s \<in> L (UPNTIMES r n)" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   822
  then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []" "length ss \<le> n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
  using Pow_cstring_nonempty by force
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
  then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []" "length vs \<le> n"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
  using IH flats_cval_nonempty by (smt order.trans) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
  then show "\<exists>v. \<Turnstile> v : UPNTIMES r n \<and> flat v = s"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   827
    using Prf.intros(7) flat_Stars by blast
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   828
next 
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   829
  case (NOT r)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   830
  then show ?case sorry
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
qed (auto intro: Prf.intros)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   832
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   833
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   834
lemma L_flat_Prf:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   835
  shows "L(r) = {flat v | v. \<Turnstile> v : r}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   836
using L_flat_Prf1 L_flat_Prf2 by blast
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   837
293
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   838
thm Prf.intros
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   839
thm Prf.cases
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   840
293
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   841
lemma
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   842
  assumes "\<Turnstile> v : (STAR r)" 
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   843
  shows "\<Turnstile> v : (FROMNTIMES r 0)"
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   844
  using assms
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   845
  apply(erule_tac Prf.cases)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   846
             apply(simp_all)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   847
  apply(case_tac vs)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   848
   apply(auto)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   849
  apply(subst append_Nil[symmetric])
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   850
   apply(rule Prf.intros)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   851
     apply(auto)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   852
  apply(simp add: Prf.intros)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   853
  done
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   854
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   855
lemma
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   856
  assumes "\<Turnstile> v : (FROMNTIMES r 0)" 
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   857
  shows "\<Turnstile> v : (STAR r)"
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   858
  using assms
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   859
  apply(erule_tac Prf.cases)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   860
             apply(simp_all)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   861
   apply(rule Prf.intros)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   862
   apply(simp)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   863
  apply(rule Prf.intros)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   864
   apply(simp)
1a4e5b94293b updated
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   865
  done
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   866
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   867
section {* Sets of Lexical Values *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   868
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   869
text {*
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   870
  Shows that lexical values are finite for a given regex and string.
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   871
*}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   872
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   873
definition
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   874
  LV :: "rexp \<Rightarrow> string \<Rightarrow> val set"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   875
where  "LV r s \<equiv> {v. \<Turnstile> v : r \<and> flat v = s}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   876
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   877
lemma LV_simps:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   878
  shows "LV ZERO s = {}"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   879
  and   "LV ONE s = (if s = [] then {Void} else {})"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
   880
  and   "LV (CH c) s = (if s = [c] then {Char c} else {})"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   881
  and   "LV (ALT r1 r2) s = Left ` LV r1 s \<union> Right ` LV r2 s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   882
unfolding LV_def
274
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   883
apply(auto intro: Prf.intros elim: Prf.cases)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   884
done
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   885
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   886
abbreviation
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   887
  "Prefixes s \<equiv> {s'. prefix s' s}"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   888
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   889
abbreviation
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   890
  "Suffixes s \<equiv> {s'. suffix s' s}"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   891
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   892
abbreviation
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   893
  "SSuffixes s \<equiv> {s'. strict_suffix s' s}"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   894
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   895
lemma Suffixes_cons [simp]:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   896
  shows "Suffixes (c # s) = Suffixes s \<union> {c # s}"
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   897
by (auto simp add: suffix_def Cons_eq_append_conv)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   898
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   899
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   900
lemma finite_Suffixes: 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   901
  shows "finite (Suffixes s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   902
by (induct s) (simp_all)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   903
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   904
lemma finite_SSuffixes: 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   905
  shows "finite (SSuffixes s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   906
proof -
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   907
  have "SSuffixes s \<subseteq> Suffixes s"
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   908
   unfolding suffix_def strict_suffix_def by auto
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   909
  then show "finite (SSuffixes s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   910
   using finite_Suffixes finite_subset by blast
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   911
qed
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   912
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   913
lemma finite_Prefixes: 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   914
  shows "finite (Prefixes s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   915
proof -
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   916
  have "finite (Suffixes (rev s))" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   917
    by (rule finite_Suffixes)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   918
  then have "finite (rev ` Suffixes (rev s))" by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   919
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   920
  have "rev ` (Suffixes (rev s)) = Prefixes s"
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   921
  unfolding suffix_def prefix_def image_def
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   922
   by (auto)(metis rev_append rev_rev_ident)+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   923
  ultimately show "finite (Prefixes s)" by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   924
qed
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   925
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   926
definition
a3134f7de065 updated
cu
parents: 275
diff changeset
   927
  "Stars_Cons V Vs \<equiv> {Stars (v # vs) | v vs. v \<in> V \<and> Stars vs \<in> Vs}"
a3134f7de065 updated
cu
parents: 275
diff changeset
   928
  
a3134f7de065 updated
cu
parents: 275
diff changeset
   929
definition
a3134f7de065 updated
cu
parents: 275
diff changeset
   930
  "Stars_Append Vs1 Vs2 \<equiv> {Stars (vs1 @ vs2) | vs1 vs2. Stars vs1 \<in> Vs1 \<and> Stars vs2 \<in> Vs2}"
a3134f7de065 updated
cu
parents: 275
diff changeset
   931
a3134f7de065 updated
cu
parents: 275
diff changeset
   932
fun Stars_Pow :: "val set \<Rightarrow> nat \<Rightarrow> val set"
a3134f7de065 updated
cu
parents: 275
diff changeset
   933
where  
a3134f7de065 updated
cu
parents: 275
diff changeset
   934
  "Stars_Pow Vs 0 = {Stars []}"
a3134f7de065 updated
cu
parents: 275
diff changeset
   935
| "Stars_Pow Vs (Suc n) = Stars_Cons Vs (Stars_Pow Vs n)"
a3134f7de065 updated
cu
parents: 275
diff changeset
   936
  
a3134f7de065 updated
cu
parents: 275
diff changeset
   937
lemma finite_Stars_Cons:
a3134f7de065 updated
cu
parents: 275
diff changeset
   938
  assumes "finite V" "finite Vs"
a3134f7de065 updated
cu
parents: 275
diff changeset
   939
  shows "finite (Stars_Cons V Vs)"
a3134f7de065 updated
cu
parents: 275
diff changeset
   940
  using assms  
a3134f7de065 updated
cu
parents: 275
diff changeset
   941
proof -
a3134f7de065 updated
cu
parents: 275
diff changeset
   942
  from assms(2) have "finite (Stars -` Vs)"
a3134f7de065 updated
cu
parents: 275
diff changeset
   943
    by(simp add: finite_vimageI inj_on_def) 
a3134f7de065 updated
cu
parents: 275
diff changeset
   944
  with assms(1) have "finite (V \<times> (Stars -` Vs))"
a3134f7de065 updated
cu
parents: 275
diff changeset
   945
    by(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
   946
  then have "finite ((\<lambda>(v, vs). Stars (v # vs)) ` (V \<times> (Stars -` Vs)))"
a3134f7de065 updated
cu
parents: 275
diff changeset
   947
    by simp
a3134f7de065 updated
cu
parents: 275
diff changeset
   948
  moreover have "Stars_Cons V Vs = (\<lambda>(v, vs). Stars (v # vs)) ` (V \<times> (Stars -` Vs))"
a3134f7de065 updated
cu
parents: 275
diff changeset
   949
    unfolding Stars_Cons_def by auto    
a3134f7de065 updated
cu
parents: 275
diff changeset
   950
  ultimately show "finite (Stars_Cons V Vs)"   
a3134f7de065 updated
cu
parents: 275
diff changeset
   951
    by simp
a3134f7de065 updated
cu
parents: 275
diff changeset
   952
qed
a3134f7de065 updated
cu
parents: 275
diff changeset
   953
a3134f7de065 updated
cu
parents: 275
diff changeset
   954
lemma finite_Stars_Append:
a3134f7de065 updated
cu
parents: 275
diff changeset
   955
  assumes "finite Vs1" "finite Vs2"
a3134f7de065 updated
cu
parents: 275
diff changeset
   956
  shows "finite (Stars_Append Vs1 Vs2)"
a3134f7de065 updated
cu
parents: 275
diff changeset
   957
  using assms  
a3134f7de065 updated
cu
parents: 275
diff changeset
   958
proof -
a3134f7de065 updated
cu
parents: 275
diff changeset
   959
  define UVs1 where "UVs1 \<equiv> Stars -` Vs1"
a3134f7de065 updated
cu
parents: 275
diff changeset
   960
  define UVs2 where "UVs2 \<equiv> Stars -` Vs2"  
a3134f7de065 updated
cu
parents: 275
diff changeset
   961
  from assms have "finite UVs1" "finite UVs2"
a3134f7de065 updated
cu
parents: 275
diff changeset
   962
    unfolding UVs1_def UVs2_def
a3134f7de065 updated
cu
parents: 275
diff changeset
   963
    by(simp_all add: finite_vimageI inj_on_def) 
a3134f7de065 updated
cu
parents: 275
diff changeset
   964
  then have "finite ((\<lambda>(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \<times> UVs2))"
a3134f7de065 updated
cu
parents: 275
diff changeset
   965
    by simp
a3134f7de065 updated
cu
parents: 275
diff changeset
   966
  moreover 
a3134f7de065 updated
cu
parents: 275
diff changeset
   967
    have "Stars_Append Vs1 Vs2 = (\<lambda>(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \<times> UVs2)"
a3134f7de065 updated
cu
parents: 275
diff changeset
   968
    unfolding Stars_Append_def UVs1_def UVs2_def by auto    
a3134f7de065 updated
cu
parents: 275
diff changeset
   969
  ultimately show "finite (Stars_Append Vs1 Vs2)"   
a3134f7de065 updated
cu
parents: 275
diff changeset
   970
    by simp
a3134f7de065 updated
cu
parents: 275
diff changeset
   971
qed 
a3134f7de065 updated
cu
parents: 275
diff changeset
   972
 
a3134f7de065 updated
cu
parents: 275
diff changeset
   973
lemma finite_Stars_Pow:
a3134f7de065 updated
cu
parents: 275
diff changeset
   974
  assumes "finite Vs"
a3134f7de065 updated
cu
parents: 275
diff changeset
   975
  shows "finite (Stars_Pow Vs n)"    
a3134f7de065 updated
cu
parents: 275
diff changeset
   976
by (induct n) (simp_all add: finite_Stars_Cons assms)
a3134f7de065 updated
cu
parents: 275
diff changeset
   977
    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   978
lemma LV_STAR_finite:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   979
  assumes "\<forall>s. finite (LV r s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   980
  shows "finite (LV (STAR r) s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   981
proof(induct s rule: length_induct)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   982
  fix s::"char list"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   983
  assume "\<forall>s'. length s' < length s \<longrightarrow> finite (LV (STAR r) s')"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   984
  then have IH: "\<forall>s' \<in> SSuffixes s. finite (LV (STAR r) s')"
279
f754a10875c7 updated for Isabelle 2017
cu
parents: 278
diff changeset
   985
    apply(auto simp add: strict_suffix_def suffix_def)
f754a10875c7 updated for Isabelle 2017
cu
parents: 278
diff changeset
   986
    by force    
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   987
  define f where "f \<equiv> \<lambda>(v, vs). Stars (v # vs)"
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
   988
  define S1 where "S1 \<equiv> \<Union>s' \<in> Prefixes s. LV r s'"
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   989
  define S2 where "S2 \<equiv> \<Union>s2 \<in> SSuffixes s. LV (STAR r) s2"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   990
  have "finite S1" using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   991
    unfolding S1_def by (simp_all add: finite_Prefixes)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   992
  moreover 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   993
  with IH have "finite S2" unfolding S2_def
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   994
    by (auto simp add: finite_SSuffixes)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   995
  ultimately 
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   996
  have "finite ({Stars []} \<union> Stars_Cons S1 S2)" 
a3134f7de065 updated
cu
parents: 275
diff changeset
   997
    by (simp add: finite_Stars_Cons)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   998
  moreover 
276
a3134f7de065 updated
cu
parents: 275
diff changeset
   999
  have "LV (STAR r) s \<subseteq> {Stars []} \<union> (Stars_Cons S1 S2)" 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1000
  unfolding S1_def S2_def f_def LV_def Stars_Cons_def
a3134f7de065 updated
cu
parents: 275
diff changeset
  1001
  unfolding prefix_def strict_suffix_def 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1002
  unfolding image_def
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1003
  apply(auto)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1004
  apply(case_tac x)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1005
  apply(auto elim: Prf_elims)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1006
  apply(erule Prf_elims)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1007
  apply(auto)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1008
  apply(case_tac vs)
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1009
  apply(auto intro: Prf.intros)  
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1010
  apply(rule exI)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1011
  apply(rule conjI)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1012
  apply(rule_tac x="flats list" in exI)
279
f754a10875c7 updated for Isabelle 2017
cu
parents: 278
diff changeset
  1013
   apply(rule conjI)
f754a10875c7 updated for Isabelle 2017
cu
parents: 278
diff changeset
  1014
  apply(simp add: suffix_def)
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1015
  apply(blast)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1016
  using Prf.intros(6) flat_Stars by blast  
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
  ultimately
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
  show "finite (LV (STAR r) s)" by (simp add: finite_subset)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
qed  
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
    
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
lemma LV_UPNTIMES_STAR:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1022
  "LV (UPNTIMES r n) s \<subseteq> LV (STAR r) s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
by(auto simp add: LV_def intro: Prf.intros elim: Prf_elims)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1024
274
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1025
lemma LV_NTIMES_3:
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1026
  shows "LV (NTIMES r (Suc n)) [] = (\<lambda>(v,vs). Stars (v#vs)) ` (LV r [] \<times> (Stars -` (LV (NTIMES r n) [])))"
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1027
unfolding LV_def
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1028
apply(auto elim!: Prf_elims simp add: image_def)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1029
apply(case_tac vs1)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1030
apply(auto)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1031
apply(case_tac vs2)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1032
apply(auto)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1033
apply(subst append.simps(1)[symmetric])
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1034
apply(rule Prf.intros)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1035
apply(auto)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1036
apply(subst append.simps(1)[symmetric])
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1037
apply(rule Prf.intros)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1038
apply(auto)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1039
  done 
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1040
    
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1041
lemma LV_FROMNTIMES_3:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1042
  shows "LV (FROMNTIMES r (Suc n)) [] = 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1043
    (\<lambda>(v,vs). Stars (v#vs)) ` (LV r [] \<times> (Stars -` (LV (FROMNTIMES r n) [])))"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1044
unfolding LV_def
a3134f7de065 updated
cu
parents: 275
diff changeset
  1045
apply(auto elim!: Prf_elims simp add: image_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1046
apply(case_tac vs1)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1047
apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1048
apply(case_tac vs2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1049
apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1050
apply(subst append.simps(1)[symmetric])
a3134f7de065 updated
cu
parents: 275
diff changeset
  1051
apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1052
     apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1053
  apply (metis le_imp_less_Suc length_greater_0_conv less_antisym list.exhaust list.set_intros(1) not_less_eq zero_le)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1054
  prefer 2
a3134f7de065 updated
cu
parents: 275
diff changeset
  1055
  using nth_mem apply blast
a3134f7de065 updated
cu
parents: 275
diff changeset
  1056
  apply(case_tac vs1)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1057
  apply (smt Groups.add_ac(2) Prf.intros(9) add.right_neutral add_Suc_right append.simps(1) insert_iff length_append list.set(2) list.size(3) list.size(4))
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1058
    apply(auto)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1059
done     
a3134f7de065 updated
cu
parents: 275
diff changeset
  1060
  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1061
lemma LV_NTIMES_4:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1062
 "LV (NTIMES r n) [] = Stars_Pow (LV r []) n" 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1063
  apply(induct n)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1064
   apply(simp add: LV_def)    
a3134f7de065 updated
cu
parents: 275
diff changeset
  1065
   apply(auto elim!: Prf_elims simp add: image_def)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1066
   apply(subst append.simps[symmetric])
a3134f7de065 updated
cu
parents: 275
diff changeset
  1067
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1068
      apply(simp_all)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1069
    apply(simp add: LV_NTIMES_3 image_def Stars_Cons_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1070
  apply blast
a3134f7de065 updated
cu
parents: 275
diff changeset
  1071
 done   
274
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1072
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1073
lemma LV_NTIMES_5:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1074
  "LV (NTIMES r n) s \<subseteq> Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (NTIMES r i) [])"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1075
apply(auto simp add: LV_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1076
apply(auto elim!: Prf_elims)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1077
  apply(auto simp add: Stars_Append_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1078
  apply(rule_tac x="vs1" in exI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1079
  apply(rule_tac x="vs2" in exI)  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1080
  apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1081
    using Prf.intros(6) apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1082
      apply(rule_tac x="length vs2" in bexI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1083
    thm Prf.intros
a3134f7de065 updated
cu
parents: 275
diff changeset
  1084
      apply(subst append.simps(1)[symmetric])
a3134f7de065 updated
cu
parents: 275
diff changeset
  1085
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1086
      apply(auto)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1087
      apply(auto)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1088
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1089
    apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1090
      done
a3134f7de065 updated
cu
parents: 275
diff changeset
  1091
      
a3134f7de065 updated
cu
parents: 275
diff changeset
  1092
lemma ttty:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1093
 "LV (FROMNTIMES r n) [] = Stars_Pow (LV r []) n" 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1094
  apply(induct n)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1095
   apply(simp add: LV_def)    
a3134f7de065 updated
cu
parents: 275
diff changeset
  1096
   apply(auto elim: Prf_elims simp add: image_def)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1097
   prefer 2
a3134f7de065 updated
cu
parents: 275
diff changeset
  1098
    apply(subst append.simps[symmetric])
a3134f7de065 updated
cu
parents: 275
diff changeset
  1099
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1100
      apply(simp_all)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1101
   apply(erule Prf_elims) 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1102
    apply(case_tac vs1)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1103
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1104
    apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1105
   apply(case_tac x)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1106
    apply(simp_all)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1107
    apply(simp add: LV_FROMNTIMES_3 image_def Stars_Cons_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1108
  apply blast
a3134f7de065 updated
cu
parents: 275
diff changeset
  1109
 done     
a3134f7de065 updated
cu
parents: 275
diff changeset
  1110
a3134f7de065 updated
cu
parents: 275
diff changeset
  1111
lemma LV_FROMNTIMES_5:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1112
  "LV (FROMNTIMES r n) s \<subseteq> Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (FROMNTIMES r i) [])"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1113
apply(auto simp add: LV_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1114
apply(auto elim!: Prf_elims)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1115
  apply(auto simp add: Stars_Append_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1116
  apply(rule_tac x="vs1" in exI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1117
  apply(rule_tac x="vs2" in exI)  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1118
  apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1119
    using Prf.intros(6) apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1120
      apply(rule_tac x="length vs2" in bexI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1121
    thm Prf.intros
a3134f7de065 updated
cu
parents: 275
diff changeset
  1122
      apply(subst append.simps(1)[symmetric])
a3134f7de065 updated
cu
parents: 275
diff changeset
  1123
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1124
      apply(auto)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1125
      apply(auto)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1126
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1127
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1128
      apply(rule_tac x="vs" in exI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1129
    apply(rule_tac x="[]" in exI) 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1130
    apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1131
    by (metis Prf.intros(9) append_Nil atMost_iff empty_iff le_imp_less_Suc less_antisym list.set(1) nth_mem zero_le)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1132
a3134f7de065 updated
cu
parents: 275
diff changeset
  1133
lemma LV_FROMNTIMES_6:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1134
  assumes "\<forall>s. finite (LV r s)"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1135
  shows "finite (LV (FROMNTIMES r n) s)"
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1136
  apply(rule finite_subset)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1137
   apply(rule LV_FROMNTIMES_5)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1138
  apply(rule finite_Stars_Append)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1139
    apply(rule LV_STAR_finite)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1140
   apply(rule assms)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1141
  apply(rule finite_UN_I)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1142
   apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1143
  by (simp add: assms finite_Stars_Pow ttty)
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1144
    
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1145
lemma LV_NMTIMES_5:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1146
  "LV (NMTIMES r n m) s \<subseteq> Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (FROMNTIMES r i) [])"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1147
apply(auto simp add: LV_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1148
apply(auto elim!: Prf_elims)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1149
  apply(auto simp add: Stars_Append_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1150
  apply(rule_tac x="vs1" in exI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1151
  apply(rule_tac x="vs2" in exI)  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1152
  apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1153
    using Prf.intros(6) apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1154
      apply(rule_tac x="length vs2" in bexI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1155
    thm Prf.intros
a3134f7de065 updated
cu
parents: 275
diff changeset
  1156
      apply(subst append.simps(1)[symmetric])
a3134f7de065 updated
cu
parents: 275
diff changeset
  1157
    apply(rule Prf.intros)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1158
      apply(auto)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1159
      apply(auto)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1160
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1161
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1162
      apply(rule_tac x="vs" in exI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1163
    apply(rule_tac x="[]" in exI) 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1164
    apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1165
    by (metis Prf.intros(9) append_Nil atMost_iff empty_iff le_imp_less_Suc less_antisym list.set(1) nth_mem zero_le)
274
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1166
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1167
lemma LV_NMTIMES_6:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1168
  assumes "\<forall>s. finite (LV r s)"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1169
  shows "finite (LV (NMTIMES r n m) s)"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1170
  apply(rule finite_subset)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1171
   apply(rule LV_NMTIMES_5)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1172
  apply(rule finite_Stars_Append)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1173
    apply(rule LV_STAR_finite)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1174
   apply(rule assms)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1175
  apply(rule finite_UN_I)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1176
   apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1177
  by (simp add: assms finite_Stars_Pow ttty)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1178
        
a3134f7de065 updated
cu
parents: 275
diff changeset
  1179
    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1180
lemma LV_finite:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1181
  shows "finite (LV r s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1182
proof(induct r arbitrary: s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1183
  case (ZERO s) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1184
  show "finite (LV ZERO s)" by (simp add: LV_simps)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1185
next
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1186
  case (ONE s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1187
  show "finite (LV ONE s)" by (simp add: LV_simps)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1188
next
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1189
  case (CH c s)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1190
  show "finite (LV (CH c) s)" by (simp add: LV_simps)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1191
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1192
  case (ALT r1 r2 s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1193
  then show "finite (LV (ALT r1 r2) s)" by (simp add: LV_simps)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1194
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1195
  case (SEQ r1 r2 s)
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1196
  define f where "f \<equiv> \<lambda>(v1, v2). Seq v1 v2"
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1197
  define S1 where "S1 \<equiv> \<Union>s' \<in> Prefixes s. LV r1 s'"
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1198
  define S2 where "S2 \<equiv> \<Union>s' \<in> Suffixes s. LV r2 s'"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1199
  have IHs: "\<And>s. finite (LV r1 s)" "\<And>s. finite (LV r2 s)" by fact+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1200
  then have "finite S1" "finite S2" unfolding S1_def S2_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1201
    by (simp_all add: finite_Prefixes finite_Suffixes)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1202
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1203
  have "LV (SEQ r1 r2) s \<subseteq> f ` (S1 \<times> S2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1204
    unfolding f_def S1_def S2_def 
275
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1205
    unfolding LV_def image_def prefix_def suffix_def
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1206
    apply (auto elim!: Prf_elims)
deea42c83c9e updated
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
  1207
    by (metis (mono_tags, lifting) mem_Collect_eq)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1208
  ultimately 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1209
  show "finite (LV (SEQ r1 r2) s)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1210
    by (simp add: finite_subset)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1211
next
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1212
  case (STAR r s)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1213
  then show "finite (LV (STAR r) s)" by (simp add: LV_STAR_finite)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1214
next 
274
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1215
  case (UPNTIMES r n s)
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1216
  have "\<And>s. finite (LV r s)" by fact
274
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1217
  then show "finite (LV (UPNTIMES r n) s)"
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1218
  by (meson LV_STAR_finite LV_UPNTIMES_STAR rev_finite_subset)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1219
next 
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1220
  case (FROMNTIMES r n s)
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1221
  have "\<And>s. finite (LV r s)" by fact
692b62426677 updated
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
  1222
  then show "finite (LV (FROMNTIMES r n) s)"
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1223
    by (simp add: LV_FROMNTIMES_6)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1224
next 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1225
  case (NTIMES r n s)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1226
  have "\<And>s. finite (LV r s)" by fact
a3134f7de065 updated
cu
parents: 275
diff changeset
  1227
  then show "finite (LV (NTIMES r n) s)"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1228
    by (metis (no_types, lifting) LV_NTIMES_4 LV_NTIMES_5 LV_STAR_finite finite_Stars_Append finite_Stars_Pow finite_UN_I finite_atMost finite_subset)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1229
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1230
  case (NMTIMES r n m s)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1231
  have "\<And>s. finite (LV r s)" by fact
a3134f7de065 updated
cu
parents: 275
diff changeset
  1232
  then show "finite (LV (NMTIMES r n m) s)"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1233
    by (simp add: LV_NMTIMES_6)    
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1234
next 
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1235
  case (NOT r s)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1236
  then show ?case sorry
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1237
qed
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1238
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1239
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1240
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1241
section {* Our POSIX Definition *}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1242
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1243
inductive 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1244
  Posix :: "string \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<in> _ \<rightarrow> _" [100, 100, 100] 100)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1245
where
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1246
  Posix_ONE: "[] \<in> ONE \<rightarrow> Void"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1247
| Posix_CHAR: "[c] \<in> (CH c) \<rightarrow> (Char c)"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1248
| Posix_ALT1: "s \<in> r1 \<rightarrow> v \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Left v)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1249
| Posix_ALT2: "\<lbrakk>s \<in> r2 \<rightarrow> v; s \<notin> L(r1)\<rbrakk> \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Right v)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1250
| Posix_SEQ: "\<lbrakk>s1 \<in> r1 \<rightarrow> v1; s2 \<in> r2 \<rightarrow> v2;
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1251
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r1 \<and> s\<^sub>4 \<in> L r2)\<rbrakk> \<Longrightarrow> 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1252
    (s1 @ s2) \<in> (SEQ r1 r2) \<rightarrow> (Seq v1 v2)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1253
| Posix_STAR1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v \<noteq> [];
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1254
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk>
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1255
    \<Longrightarrow> (s1 @ s2) \<in> STAR r \<rightarrow> Stars (v # vs)"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1256
| Posix_STAR2: "[] \<in> STAR r \<rightarrow> Stars []"
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1257
| Posix_NTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> NTIMES r (n - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n;
a3134f7de065 updated
cu
parents: 275
diff changeset
  1258
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))\<rbrakk>
a3134f7de065 updated
cu
parents: 275
diff changeset
  1259
    \<Longrightarrow> (s1 @ s2) \<in> NTIMES r n \<rightarrow> Stars (v # vs)"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1260
| Posix_NTIMES2: "\<lbrakk>\<forall>v \<in> set vs. [] \<in> r \<rightarrow> v; length vs = n\<rbrakk>
a3134f7de065 updated
cu
parents: 275
diff changeset
  1261
    \<Longrightarrow> [] \<in> NTIMES r n \<rightarrow> Stars vs"  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1262
| Posix_UPNTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> UPNTIMES r (n - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n;
a3134f7de065 updated
cu
parents: 275
diff changeset
  1263
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (n - 1)))\<rbrakk>
a3134f7de065 updated
cu
parents: 275
diff changeset
  1264
    \<Longrightarrow> (s1 @ s2) \<in> UPNTIMES r n \<rightarrow> Stars (v # vs)"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1265
| Posix_UPNTIMES2: "[] \<in> UPNTIMES r n \<rightarrow> Stars []"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1266
| Posix_FROMNTIMES2: "\<lbrakk>\<forall>v \<in> set vs. [] \<in> r \<rightarrow> v; length vs = n\<rbrakk>
a3134f7de065 updated
cu
parents: 275
diff changeset
  1267
    \<Longrightarrow> [] \<in> FROMNTIMES r n \<rightarrow> Stars vs"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1268
| Posix_FROMNTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> FROMNTIMES r (n - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n;
a3134f7de065 updated
cu
parents: 275
diff changeset
  1269
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (FROMNTIMES r (n - 1)))\<rbrakk>
a3134f7de065 updated
cu
parents: 275
diff changeset
  1270
    \<Longrightarrow> (s1 @ s2) \<in> FROMNTIMES r n \<rightarrow> Stars (v # vs)"  
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1271
| Posix_FROMNTIMES3: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v \<noteq> [];
42268a284ea6 updated
cu
parents: 276
diff changeset
  1272
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk>
42268a284ea6 updated
cu
parents: 276
diff changeset
  1273
    \<Longrightarrow> (s1 @ s2) \<in> FROMNTIMES r 0 \<rightarrow> Stars (v # vs)"  
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1274
| Posix_NMTIMES2: "\<lbrakk>\<forall>v \<in> set vs. [] \<in> r \<rightarrow> v; length vs = n; n \<le> m\<rbrakk>
a3134f7de065 updated
cu
parents: 275
diff changeset
  1275
    \<Longrightarrow> [] \<in> NMTIMES r n m \<rightarrow> Stars vs"  
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1276
| Posix_NMTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> NMTIMES r (n - 1) (m - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n; n \<le> m;
424bdcd01016 updated
cu
parents: 277
diff changeset
  1277
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (NMTIMES r (n - 1) (m - 1)))\<rbrakk>
424bdcd01016 updated
cu
parents: 277
diff changeset
  1278
    \<Longrightarrow> (s1 @ s2) \<in> NMTIMES r n m \<rightarrow> Stars (v # vs)"  
424bdcd01016 updated
cu
parents: 277
diff changeset
  1279
| Posix_NMTIMES3: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> UPNTIMES r (m - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < m;
424bdcd01016 updated
cu
parents: 277
diff changeset
  1280
    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (m - 1)))\<rbrakk>
424bdcd01016 updated
cu
parents: 277
diff changeset
  1281
    \<Longrightarrow> (s1 @ s2) \<in> NMTIMES r 0 m \<rightarrow> Stars (v # vs)"    
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1282
  
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1283
inductive_cases Posix_elims:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1284
  "s \<in> ZERO \<rightarrow> v"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1285
  "s \<in> ONE \<rightarrow> v"
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1286
  "s \<in> CH c \<rightarrow> v"
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1287
  "s \<in> ALT r1 r2 \<rightarrow> v"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1288
  "s \<in> SEQ r1 r2 \<rightarrow> v"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1289
  "s \<in> STAR r \<rightarrow> v"
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1290
  "s \<in> NTIMES r n \<rightarrow> v"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1291
  "s \<in> UPNTIMES r n \<rightarrow> v"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1292
  "s \<in> FROMNTIMES r n \<rightarrow> v"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1293
  "s \<in> NMTIMES r n m \<rightarrow> v"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1294
  
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1295
lemma Posix1:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1296
  assumes "s \<in> r \<rightarrow> v"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1297
  shows "s \<in> L r" "flat v = s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1298
using assms
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1299
  apply(induct s r v rule: Posix.induct)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1300
                    apply(auto simp add: Sequ_def)[18]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1301
            apply(case_tac n)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1302
             apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1303
  apply(simp add: Sequ_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1304
            apply(auto)[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1305
           apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1306
  apply(clarify)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1307
  apply(rule_tac x="Suc x" in bexI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1308
  apply(simp add: Sequ_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1309
            apply(auto)[5]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1310
  using nth_mem nullable.simps(9) nullable_correctness apply auto[1]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1311
  apply simp
a3134f7de065 updated
cu
parents: 275
diff changeset
  1312
       apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1313
       apply(clarify)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1314
       apply(rule_tac x="Suc x" in bexI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1315
        apply(simp add: Sequ_def)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1316
          apply(auto)[3]
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1317
    defer
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1318
     apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1319
  apply fastforce
a3134f7de065 updated
cu
parents: 275
diff changeset
  1320
    apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1321
   apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1322
    apply(clarify)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1323
   apply(rule_tac x="Suc x" in bexI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1324
    apply(auto simp add: Sequ_def)[2]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1325
   apply(simp)
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1326
    apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1327
    apply(clarify)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1328
     apply(rule_tac x="Suc x" in bexI)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1329
    apply(auto simp add: Sequ_def)[2]
424bdcd01016 updated
cu
parents: 277
diff changeset
  1330
   apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1331
  apply(simp add: Star.step Star_Pow)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1332
done  
424bdcd01016 updated
cu
parents: 277
diff changeset
  1333
    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1334
text {*
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1335
  Our Posix definition determines a unique value.
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1336
*}
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1337
  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1338
lemma List_eq_zipI:
a3134f7de065 updated
cu
parents: 275
diff changeset
  1339
  assumes "\<forall>(v1, v2) \<in> set (zip vs1 vs2). v1 = v2" 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1340
  and "length vs1 = length vs2"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1341
  shows "vs1 = vs2"  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1342
 using assms
a3134f7de065 updated
cu
parents: 275
diff changeset
  1343
  apply(induct vs1 arbitrary: vs2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1344
   apply(case_tac vs2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1345
   apply(simp)    
a3134f7de065 updated
cu
parents: 275
diff changeset
  1346
   apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1347
   apply(case_tac vs2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1348
   apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1349
  apply(simp)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1350
done    
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1351
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1352
lemma Posix_determ:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1353
  assumes "s \<in> r \<rightarrow> v1" "s \<in> r \<rightarrow> v2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1354
  shows "v1 = v2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1355
using assms
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1356
proof (induct s r v1 arbitrary: v2 rule: Posix.induct)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1357
  case (Posix_ONE v2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1358
  have "[] \<in> ONE \<rightarrow> v2" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1359
  then show "Void = v2" by cases auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1360
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1361
  case (Posix_CHAR c v2)
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1362
  have "[c] \<in> CH c \<rightarrow> v2" by fact
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1363
  then show "Char c = v2" by cases auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1364
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1365
  case (Posix_ALT1 s r1 v r2 v2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1366
  have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1367
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1368
  have "s \<in> r1 \<rightarrow> v" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1369
  then have "s \<in> L r1" by (simp add: Posix1)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1370
  ultimately obtain v' where eq: "v2 = Left v'" "s \<in> r1 \<rightarrow> v'" by cases auto 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1371
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1372
  have IH: "\<And>v2. s \<in> r1 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1373
  ultimately have "v = v'" by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1374
  then show "Left v = v2" using eq by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1375
next 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1376
  case (Posix_ALT2 s r2 v r1 v2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1377
  have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1378
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1379
  have "s \<notin> L r1" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1380
  ultimately obtain v' where eq: "v2 = Right v'" "s \<in> r2 \<rightarrow> v'" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1381
    by cases (auto simp add: Posix1) 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1382
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1383
  have IH: "\<And>v2. s \<in> r2 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1384
  ultimately have "v = v'" by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1385
  then show "Right v = v2" using eq by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1386
next
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1387
  case (Posix_SEQ s1 r1 v1 s2 r2 v2 v')
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1388
  have "(s1 @ s2) \<in> SEQ r1 r2 \<rightarrow> v'" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1389
       "s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1390
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by fact+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1391
  then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \<in> r1 \<rightarrow> v1'" "s2 \<in> r2 \<rightarrow> v2'"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1392
  apply(cases) apply (auto simp add: append_eq_append_conv2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1393
  using Posix1(1) by fastforce+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1394
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1395
  have IHs: "\<And>v1'. s1 \<in> r1 \<rightarrow> v1' \<Longrightarrow> v1 = v1'"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1396
            "\<And>v2'. s2 \<in> r2 \<rightarrow> v2' \<Longrightarrow> v2 = v2'" by fact+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1397
  ultimately show "Seq v1 v2 = v'" by simp
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1398
next
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1399
  case (Posix_STAR1 s1 r v s2 vs v2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1400
  have "(s1 @ s2) \<in> STAR r \<rightarrow> v2" 
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1401
       "s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" "flat v \<noteq> []"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1402
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1403
  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (Stars vs')"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1404
  apply(cases) apply (auto simp add: append_eq_append_conv2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1405
  using Posix1(1) apply fastforce
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1406
  apply (metis Posix1(1) Posix_STAR1.hyps(6) append_Nil append_Nil2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1407
  using Posix1(2) by blast
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1408
  moreover
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1409
  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1410
            "\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1411
  ultimately show "Stars (v # vs) = v2" by auto
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1412
next
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1413
  case (Posix_STAR2 r v2)
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1414
  have "[] \<in> STAR r \<rightarrow> v2" by fact
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1415
  then show "Stars [] = v2" by cases (auto simp add: Posix1)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1416
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1417
  case (Posix_NTIMES2 vs r n v2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1418
  then show "Stars vs = v2"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1419
    apply(erule_tac Posix_elims)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1420
     apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1421
     apply (simp add: Posix1(2))
a3134f7de065 updated
cu
parents: 275
diff changeset
  1422
    apply(rule List_eq_zipI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1423
     apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1424
    by (meson in_set_zipE)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1425
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1426
  case (Posix_NTIMES1 s1 r v s2 n vs v2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1427
  have "(s1 @ s2) \<in> NTIMES r n \<rightarrow> v2" 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1428
       "s1 \<in> r \<rightarrow> v" "s2 \<in> NTIMES r (n - 1) \<rightarrow> Stars vs" "flat v \<noteq> []"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1429
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1 )))" by fact+
a3134f7de065 updated
cu
parents: 275
diff changeset
  1430
  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (NTIMES r (n - 1)) \<rightarrow> (Stars vs')"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1431
  apply(cases) apply (auto simp add: append_eq_append_conv2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1432
    using Posix1(1) apply fastforce
a3134f7de065 updated
cu
parents: 275
diff changeset
  1433
    apply (metis One_nat_def Posix1(1) Posix_NTIMES1.hyps(7) append.right_neutral append_self_conv2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1434
  using Posix1(2) by blast
a3134f7de065 updated
cu
parents: 275
diff changeset
  1435
  moreover
a3134f7de065 updated
cu
parents: 275
diff changeset
  1436
  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1437
            "\<And>v2. s2 \<in> NTIMES r (n - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
a3134f7de065 updated
cu
parents: 275
diff changeset
  1438
  ultimately show "Stars (v # vs) = v2" by auto
a3134f7de065 updated
cu
parents: 275
diff changeset
  1439
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1440
  case (Posix_UPNTIMES1 s1 r v s2 n vs v2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1441
  have "(s1 @ s2) \<in> UPNTIMES r n \<rightarrow> v2" 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1442
       "s1 \<in> r \<rightarrow> v" "s2 \<in> UPNTIMES r (n - 1) \<rightarrow> Stars vs" "flat v \<noteq> []"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1443
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (n - 1 )))" by fact+
a3134f7de065 updated
cu
parents: 275
diff changeset
  1444
  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (UPNTIMES r (n - 1)) \<rightarrow> (Stars vs')"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1445
  apply(cases) apply (auto simp add: append_eq_append_conv2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1446
    using Posix1(1) apply fastforce
a3134f7de065 updated
cu
parents: 275
diff changeset
  1447
    apply (metis One_nat_def Posix1(1) Posix_UPNTIMES1.hyps(7) append.right_neutral append_self_conv2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1448
  using Posix1(2) by blast
a3134f7de065 updated
cu
parents: 275
diff changeset
  1449
  moreover
a3134f7de065 updated
cu
parents: 275
diff changeset
  1450
  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1451
            "\<And>v2. s2 \<in> UPNTIMES r (n - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
a3134f7de065 updated
cu
parents: 275
diff changeset
  1452
  ultimately show "Stars (v # vs) = v2" by auto
a3134f7de065 updated
cu
parents: 275
diff changeset
  1453
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1454
  case (Posix_UPNTIMES2 r n v2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1455
  then show "Stars [] = v2"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1456
    apply(erule_tac Posix_elims)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1457
     apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1458
    by (simp add: Posix1(2))
a3134f7de065 updated
cu
parents: 275
diff changeset
  1459
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1460
  case (Posix_FROMNTIMES1 s1 r v s2 n vs v2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1461
  have "(s1 @ s2) \<in> FROMNTIMES r n \<rightarrow> v2" 
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1462
       "s1 \<in> r \<rightarrow> v" "s2 \<in> FROMNTIMES r (n - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" "0 < n"
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1463
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (FROMNTIMES r (n - 1 )))" by fact+
a3134f7de065 updated
cu
parents: 275
diff changeset
  1464
  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (FROMNTIMES r (n - 1)) \<rightarrow> (Stars vs')"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1465
  apply(cases) apply (auto simp add: append_eq_append_conv2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1466
    using Posix1(1) Posix1(2) apply blast 
a3134f7de065 updated
cu
parents: 275
diff changeset
  1467
     apply(case_tac n)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1468
      apply(simp)
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1469
     apply(simp)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1470
    apply (smt (verit, ccfv_threshold) L.simps(9) Posix1(1) UN_E append_eq_append_conv2)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1471
    by (metis One_nat_def Posix1(1) Posix_FROMNTIMES1.hyps(7) append_Nil2 append_self_conv2)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1472
  moreover
a3134f7de065 updated
cu
parents: 275
diff changeset
  1473
  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1474
            "\<And>v2. s2 \<in> FROMNTIMES r (n - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
a3134f7de065 updated
cu
parents: 275
diff changeset
  1475
  ultimately show "Stars (v # vs) = v2" by auto    
a3134f7de065 updated
cu
parents: 275
diff changeset
  1476
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1477
  case (Posix_FROMNTIMES2 vs r n v2)  
a3134f7de065 updated
cu
parents: 275
diff changeset
  1478
  then show "Stars vs = v2"
a3134f7de065 updated
cu
parents: 275
diff changeset
  1479
    apply(erule_tac Posix_elims)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1480
     apply(auto)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1481
    apply(rule List_eq_zipI)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1482
     apply(auto)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1483
      apply(meson in_set_zipE)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1484
     apply (simp add: Posix1(2))
42268a284ea6 updated
cu
parents: 276
diff changeset
  1485
    using Posix1(2) by blast
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1486
next
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1487
  case (Posix_FROMNTIMES3 s1 r v s2 vs v2)  
42268a284ea6 updated
cu
parents: 276
diff changeset
  1488
    have "(s1 @ s2) \<in> FROMNTIMES r 0 \<rightarrow> v2" 
42268a284ea6 updated
cu
parents: 276
diff changeset
  1489
       "s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" "flat v \<noteq> []"
42268a284ea6 updated
cu
parents: 276
diff changeset
  1490
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact+
42268a284ea6 updated
cu
parents: 276
diff changeset
  1491
  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (Stars vs')"
42268a284ea6 updated
cu
parents: 276
diff changeset
  1492
  apply(cases) apply (auto simp add: append_eq_append_conv2)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1493
    using Posix1(2) apply fastforce
42268a284ea6 updated
cu
parents: 276
diff changeset
  1494
    using Posix1(1) apply fastforce
42268a284ea6 updated
cu
parents: 276
diff changeset
  1495
    by (metis Posix1(1) Posix_FROMNTIMES3.hyps(6) append.right_neutral append_Nil)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1496
  moreover
42268a284ea6 updated
cu
parents: 276
diff changeset
  1497
  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
42268a284ea6 updated
cu
parents: 276
diff changeset
  1498
            "\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
42268a284ea6 updated
cu
parents: 276
diff changeset
  1499
  ultimately show "Stars (v # vs) = v2" by auto     
42268a284ea6 updated
cu
parents: 276
diff changeset
  1500
next    
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1501
  case (Posix_NMTIMES1 s1 r v s2 n m vs v2)
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1502
  have "(s1 @ s2) \<in> NMTIMES r n m \<rightarrow> v2" 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1503
       "s1 \<in> r \<rightarrow> v" "s2 \<in> NMTIMES r (n - 1) (m - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1504
       "0 < n" "n \<le> m"
424bdcd01016 updated
cu
parents: 277
diff changeset
  1505
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NMTIMES r (n - 1) (m - 1)))" by fact+
424bdcd01016 updated
cu
parents: 277
diff changeset
  1506
  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1507
    "s2 \<in> (NMTIMES r (n - 1) (m - 1)) \<rightarrow> (Stars vs')"
424bdcd01016 updated
cu
parents: 277
diff changeset
  1508
  apply(cases) apply (auto simp add: append_eq_append_conv2)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1509
    using Posix1(1) Posix1(2) apply blast 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1510
     apply(case_tac n)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1511
      apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1512
     apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1513
       apply(case_tac m)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1514
      apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1515
     apply(simp)
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1516
    apply (metis L.simps(10) Posix1(1) UN_E append.right_neutral append_Nil)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 359
diff changeset
  1517
    by (metis One_nat_def Posix1(1) Posix_NMTIMES1.hyps(8) append_Nil self_append_conv)
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1518
  moreover
424bdcd01016 updated
cu
parents: 277
diff changeset
  1519
  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
424bdcd01016 updated
cu
parents: 277
diff changeset
  1520
            "\<And>v2. s2 \<in> NMTIMES r (n - 1) (m - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
424bdcd01016 updated
cu
parents: 277
diff changeset
  1521
  ultimately show "Stars (v # vs) = v2" by auto     
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1522
next
a3134f7de065 updated
cu
parents: 275
diff changeset
  1523
  case (Posix_NMTIMES2 vs r n m v2)
a3134f7de065 updated
cu
parents: 275
diff changeset
  1524
  then show "Stars vs = v2"
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1525
    apply(erule_tac Posix_elims)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1526
      apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1527
      apply(rule List_eq_zipI)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1528
       apply(auto)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1529
      apply (meson in_set_zipE)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1530
    apply (simp add: Posix1(2))
424bdcd01016 updated
cu
parents: 277
diff changeset
  1531
    apply(erule_tac Posix_elims)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1532
     apply(auto)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1533
    apply (simp add: Posix1(2))+
424bdcd01016 updated
cu
parents: 277
diff changeset
  1534
    done  
424bdcd01016 updated
cu
parents: 277
diff changeset
  1535
next
424bdcd01016 updated
cu
parents: 277
diff changeset
  1536
  case (Posix_NMTIMES3 s1 r v s2 m vs v2)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1537
   have "(s1 @ s2) \<in> NMTIMES r 0 m \<rightarrow> v2" 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1538
       "s1 \<in> r \<rightarrow> v" "s2 \<in> UPNTIMES r (m - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" "0 < m"
424bdcd01016 updated
cu
parents: 277
diff changeset
  1539
       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (m - 1 )))" by fact+
424bdcd01016 updated
cu
parents: 277
diff changeset
  1540
  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (UPNTIMES r (m - 1)) \<rightarrow> (Stars vs')"
424bdcd01016 updated
cu
parents: 277
diff changeset
  1541
    apply(cases) apply (auto simp add: append_eq_append_conv2)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1542
    using Posix1(2) apply blast
424bdcd01016 updated
cu
parents: 277
diff changeset
  1543
    apply (smt L.simps(7) Posix1(1) UN_E append_eq_append_conv2)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1544
    by (metis One_nat_def Posix1(1) Posix_NMTIMES3.hyps(7) append.right_neutral append_Nil)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1545
  moreover
424bdcd01016 updated
cu
parents: 277
diff changeset
  1546
  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
424bdcd01016 updated
cu
parents: 277
diff changeset
  1547
            "\<And>v2. s2 \<in> UPNTIMES r (m - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
424bdcd01016 updated
cu
parents: 277
diff changeset
  1548
  ultimately show "Stars (v # vs) = v2" by auto  
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1549
qed
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1550
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1551
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1552
text {*
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1553
  Our POSIX value is a lexical value.
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1554
*}
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1555
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1556
lemma Posix_LV:
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1557
  assumes "s \<in> r \<rightarrow> v"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1558
  shows "v \<in> LV r s"
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1559
using assms unfolding LV_def
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1560
apply(induct rule: Posix.induct)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1561
            apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)[7]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1562
     defer
a3134f7de065 updated
cu
parents: 275
diff changeset
  1563
  defer
a3134f7de065 updated
cu
parents: 275
diff changeset
  1564
     apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)[2]
a3134f7de065 updated
cu
parents: 275
diff changeset
  1565
  apply (metis (mono_tags, lifting) Prf.intros(9) append_Nil empty_iff flat_Stars flats_empty list.set(1) mem_Collect_eq)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1566
     apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1567
     apply(case_tac n)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1568
      apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1569
     apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1570
     apply(erule Prf_elims)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1571
      apply(simp)
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1572
  apply(subst append.simps(2)[symmetric])
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1573
      apply(rule Prf.intros) 
42268a284ea6 updated
cu
parents: 276
diff changeset
  1574
        apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1575
       apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1576
      apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1577
     apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1578
     apply(rule Prf.intros)  
42268a284ea6 updated
cu
parents: 276
diff changeset
  1579
      apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1580
     apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1581
    apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1582
   apply(erule Prf_elims)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1583
      apply(simp)
42268a284ea6 updated
cu
parents: 276
diff changeset
  1584
  apply(rule Prf.intros)  
42268a284ea6 updated
cu
parents: 276
diff changeset
  1585
       apply(simp)
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1586
     apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1587
  (* NTIMES *)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1588
   prefer 4
424bdcd01016 updated
cu
parents: 277
diff changeset
  1589
   apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1590
   apply(case_tac n)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1591
    apply(simp)
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1592
   apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1593
   apply(clarify)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1594
   apply(rotate_tac 5)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1595
   apply(erule Prf_elims)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1596
   apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1597
  apply(subst append.simps(2)[symmetric])
424bdcd01016 updated
cu
parents: 277
diff changeset
  1598
      apply(rule Prf.intros) 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1599
        apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1600
       apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1601
   apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1602
  prefer 4
424bdcd01016 updated
cu
parents: 277
diff changeset
  1603
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1604
  apply (metis Prf.intros(8) length_removeAll_less less_irrefl_nat removeAll.simps(1) self_append_conv2)
277
42268a284ea6 updated
cu
parents: 276
diff changeset
  1605
  (* NMTIMES *)
278
424bdcd01016 updated
cu
parents: 277
diff changeset
  1606
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1607
  apply (metis Prf.intros(11) append_Nil empty_iff list.set(1))
424bdcd01016 updated
cu
parents: 277
diff changeset
  1608
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1609
  apply(rotate_tac 6)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1610
  apply(erule Prf_elims)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1611
   apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1612
  apply(subst append.simps(2)[symmetric])
424bdcd01016 updated
cu
parents: 277
diff changeset
  1613
      apply(rule Prf.intros) 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1614
        apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1615
       apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1616
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1617
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1618
  apply(rule Prf.intros) 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1619
        apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1620
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1621
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1622
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1623
  apply(rotate_tac 6)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1624
  apply(erule Prf_elims)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1625
   apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1626
      apply(rule Prf.intros) 
424bdcd01016 updated
cu
parents: 277
diff changeset
  1627
        apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1628
       apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1629
  apply(simp)
424bdcd01016 updated
cu
parents: 277
diff changeset
  1630
done    
276
a3134f7de065 updated
cu
parents: 275
diff changeset
  1631
  
273
e85099ac4c6c updated
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1632
end