AFP-Submission/Regular_Exp.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 03 Jun 2016 11:07:10 +0100
changeset 193 1fd7388360b6
parent 191 6bb15b8e6301
permissions -rw-r--r--
typos
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
191
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
(*  Author: Tobias Nipkow *)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
section "Regular expressions"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
theory Regular_Exp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
imports Regular_Set
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
begin
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
datatype (atoms: 'a) rexp =
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
  is_Zero: Zero |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
  is_One: One |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
  Atom 'a |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
  Plus "('a rexp)" "('a rexp)" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
  Times "('a rexp)" "('a rexp)" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
  Star "('a rexp)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
primrec lang :: "'a rexp => 'a lang" where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
"lang Zero = {}" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
"lang One = {[]}" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
"lang (Atom a) = {[a]}" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
"lang (Plus r s) = (lang r) Un (lang s)" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
"lang (Times r s) = conc (lang r) (lang s)" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
"lang (Star r) = star(lang r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
primrec nullable :: "'a rexp \<Rightarrow> bool" where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
"nullable Zero = False" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
"nullable One = True" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
"nullable (Atom c) = False" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
"nullable (Plus r1 r2) = (nullable r1 \<or> nullable r2)" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
"nullable (Times r1 r2) = (nullable r1 \<and> nullable r2)" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
"nullable (Star r) = True"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
lemma nullable_iff: "nullable r \<longleftrightarrow> [] \<in> lang r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
by (induct r) (auto simp add: conc_def split: if_splits)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
text{* Composition on rhs usually complicates matters: *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
lemma map_map_rexp:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
  "map_rexp f (map_rexp g r) = map_rexp (\<lambda>r. f (g r)) r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
  unfolding rexp.map_comp o_def ..
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
lemma map_rexp_ident[simp]: "map_rexp (\<lambda>x. x) = (\<lambda>r. r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
  unfolding id_def[symmetric] fun_eq_iff rexp.map_id id_apply by (intro allI refl)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
lemma atoms_lang: "w : lang r \<Longrightarrow> set w \<subseteq> atoms r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
proof(induction r arbitrary: w)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
  case Times thus ?case by fastforce
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
  case Star thus ?case by (fastforce simp add: star_conv_concat)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
qed auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
lemma lang_eq_ext: "(lang r = lang s) =
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
  (\<forall>w \<in> lists(atoms r \<union> atoms s). w \<in> lang r \<longleftrightarrow> w \<in> lang s)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
  by (auto simp: atoms_lang[unfolded subset_iff])
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
lemma lang_eq_ext_Nil_fold_Deriv:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
  fixes r s
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
  defines "\<BB> \<equiv> {(fold Deriv w (lang r), fold Deriv w (lang s))| w. w\<in>lists (atoms r \<union> atoms s)}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
  shows "lang r = lang s \<longleftrightarrow> (\<forall>(K, L) \<in> \<BB>. [] \<in> K \<longleftrightarrow> [] \<in> L)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
  unfolding lang_eq_ext \<BB>_def by (subst (1 2) in_fold_Deriv[of "[]", simplified, symmetric]) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
subsection {* Term ordering *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
instantiation rexp :: (order) "{order}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
begin
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
fun le_rexp :: "('a::order) rexp \<Rightarrow> ('a::order) rexp \<Rightarrow> bool"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
  "le_rexp Zero _ = True"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
| "le_rexp _ Zero = False"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
| "le_rexp One _ = True"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
| "le_rexp _ One = False"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
| "le_rexp (Atom a) (Atom b) = (a <= b)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
| "le_rexp (Atom _) _ = True"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
| "le_rexp _ (Atom _) = False"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
| "le_rexp (Star r) (Star s) = le_rexp r s"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
| "le_rexp (Star _) _ = True"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
| "le_rexp _ (Star _) = False"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
| "le_rexp (Plus r r') (Plus s s') =
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
    (if r = s then le_rexp r' s' else le_rexp r s)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
| "le_rexp (Plus _ _) _ = True"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
| "le_rexp _ (Plus _ _) = False"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
| "le_rexp (Times r r') (Times s s') =
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
    (if r = s then le_rexp r' s' else le_rexp r s)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
(* The class instance stuff is by Dmitriy Traytel *)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
definition less_eq_rexp where "r \<le> s \<equiv> le_rexp r s"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
definition less_rexp where "r < s \<equiv> le_rexp r s \<and> r \<noteq> s"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
lemma le_rexp_Zero: "le_rexp r Zero \<Longrightarrow> r = Zero"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
by (induction r) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
lemma le_rexp_refl: "le_rexp r r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
by (induction r) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
lemma le_rexp_antisym: "\<lbrakk>le_rexp r s; le_rexp s r\<rbrakk> \<Longrightarrow> r = s"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
by (induction r s rule: le_rexp.induct) (auto dest: le_rexp_Zero)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
lemma le_rexp_trans: "\<lbrakk>le_rexp r s; le_rexp s t\<rbrakk> \<Longrightarrow> le_rexp r t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
proof (induction r s arbitrary: t rule: le_rexp.induct)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
  fix v t assume "le_rexp (Atom v) t" thus "le_rexp One t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
  fix s1 s2 t assume "le_rexp (Plus s1 s2) t" thus "le_rexp One t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
  fix s1 s2 t assume "le_rexp (Times s1 s2) t" thus "le_rexp One t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
  fix s t assume "le_rexp (Star s) t" thus "le_rexp One t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
  fix v u t assume "le_rexp (Atom v) (Atom u)" "le_rexp (Atom u) t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
  thus "le_rexp (Atom v) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
  fix v s1 s2 t assume "le_rexp (Plus s1 s2) t" thus "le_rexp (Atom v) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
  fix v s1 s2 t assume "le_rexp (Times s1 s2) t" thus "le_rexp (Atom v) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
  fix v s t assume "le_rexp (Star s) t" thus "le_rexp (Atom v) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
  fix r s t
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
  assume IH: "\<And>t. le_rexp r s \<Longrightarrow> le_rexp s t \<Longrightarrow> le_rexp r t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
    and "le_rexp (Star r) (Star s)" "le_rexp (Star s) t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
  thus "le_rexp (Star r) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
  fix r s1 s2 t assume "le_rexp (Plus s1 s2) t" thus "le_rexp (Star r) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
  fix r s1 s2 t assume "le_rexp (Times s1 s2) t" thus "le_rexp (Star r) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
  fix r1 r2 s1 s2 t
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
  assume "\<And>t. r1 = s1 \<Longrightarrow> le_rexp r2 s2 \<Longrightarrow> le_rexp s2 t \<Longrightarrow> le_rexp r2 t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
         "\<And>t. r1 \<noteq> s1 \<Longrightarrow> le_rexp r1 s1 \<Longrightarrow> le_rexp s1 t \<Longrightarrow> le_rexp r1 t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
         "le_rexp (Plus r1 r2) (Plus s1 s2)" "le_rexp (Plus s1 s2) t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
  thus "le_rexp (Plus r1 r2) t" by (cases t) (auto split: split_if_asm intro: le_rexp_antisym)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
  fix r1 r2 s1 s2 t assume "le_rexp (Times s1 s2) t" thus "le_rexp (Plus r1 r2) t" by (cases t) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
  fix r1 r2 s1 s2 t
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
  assume "\<And>t. r1 = s1 \<Longrightarrow> le_rexp r2 s2 \<Longrightarrow> le_rexp s2 t \<Longrightarrow> le_rexp r2 t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
         "\<And>t. r1 \<noteq> s1 \<Longrightarrow> le_rexp r1 s1 \<Longrightarrow> le_rexp s1 t \<Longrightarrow> le_rexp r1 t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
         "le_rexp (Times r1 r2) (Times s1 s2)" "le_rexp (Times s1 s2) t"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
  thus "le_rexp (Times r1 r2) t" by (cases t) (auto split: split_if_asm intro: le_rexp_antisym)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
qed auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
instance proof
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
qed (auto simp add: less_eq_rexp_def less_rexp_def
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
       intro: le_rexp_refl le_rexp_antisym le_rexp_trans)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
end
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
instantiation rexp :: (linorder) "{linorder}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
begin
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
lemma le_rexp_total: "le_rexp (r :: 'a :: linorder rexp) s \<or> le_rexp s r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
by (induction r s rule: le_rexp.induct) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
instance proof
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
qed (unfold less_eq_rexp_def less_rexp_def, rule le_rexp_total)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
end
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
end