diff -r e752d84225ec -r 8bb064045b4e thys/RegLangs.thy --- a/thys/RegLangs.thy Mon Feb 22 03:22:26 2021 +0000 +++ b/thys/RegLangs.thy Thu Feb 25 22:46:58 2021 +0000 @@ -1,16 +1,15 @@ - theory RegLangs imports Main "~~/src/HOL/Library/Sublist" begin -section {* Sequential Composition of Languages *} +section \Sequential Composition of Languages\ definition Sequ :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) where "A ;; B = {s1 @ s2 | s1 s2. s1 \ A \ s2 \ B}" -text {* Two Simple Properties about Sequential Composition *} +text \Two Simple Properties about Sequential Composition\ lemma Sequ_empty_string [simp]: shows "A ;; {[]} = A" @@ -22,22 +21,8 @@ and "{} ;; A = {}" by (simp_all add: Sequ_def) -lemma - shows "(A \ B) ;; C = (A ;; C) \ (B ;; C)" - apply(auto simp add: Sequ_def) - done -lemma - shows "(A \ B) ;; C \ (A ;; C) \ (B ;; C)" - apply(auto simp add: Sequ_def) - done - -lemma - shows "(A ;; C) \ (B ;; C) \ (A \ B) ;; C" - apply(auto simp add: Sequ_def) - oops - -section {* Semantic Derivative (Left Quotient) of Languages *} +section \Semantic Derivative (Left Quotient) of Languages\ definition Der :: "char \ string set \ string set" @@ -75,7 +60,7 @@ by (auto simp add: Cons_eq_append_conv) -section {* Kleene Star for Languages *} +section \Kleene Star for Languages\ inductive_set Star :: "string set \ string set" ("_\" [101] 102) @@ -104,7 +89,7 @@ by(auto simp add: Star_decomp) -lemma Der_star [simp]: +lemma Der_star[simp]: shows "Der c (A\) = (Der c A) ;; A\" proof - have "Der c (A\) = Der c ({[]} \ A ;; A\)" @@ -134,37 +119,37 @@ -section {* Regular Expressions *} +section \Regular Expressions\ datatype rexp = ZERO | ONE -| CHAR char +| CH char | SEQ rexp rexp | ALT rexp rexp | STAR rexp -section {* Semantics of Regular Expressions *} +section \Semantics of Regular Expressions\ fun L :: "rexp \ string set" where "L (ZERO) = {}" | "L (ONE) = {[]}" -| "L (CHAR c) = {[c]}" +| "L (CH c) = {[c]}" | "L (SEQ r1 r2) = (L r1) ;; (L r2)" | "L (ALT r1 r2) = (L r1) \ (L r2)" | "L (STAR r) = (L r)\" -section {* Nullable, Derivatives *} +section \Nullable, Derivatives\ fun nullable :: "rexp \ bool" where "nullable (ZERO) = False" | "nullable (ONE) = True" -| "nullable (CHAR c) = False" +| "nullable (CH c) = False" | "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" | "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" | "nullable (STAR r) = True" @@ -175,7 +160,7 @@ where "der c (ZERO) = ZERO" | "der c (ONE) = ZERO" -| "der c (CHAR d) = (if c = d then ONE else ZERO)" +| "der c (CH d) = (if c = d then ONE else ZERO)" | "der c (ALT r1 r2) = ALT (der c r1) (der c r2)" | "der c (SEQ r1 r2) = (if nullable r1