diff -r e752d84225ec -r 8bb064045b4e thys/Spec.thy --- a/thys/Spec.thy Mon Feb 22 03:22:26 2021 +0000 +++ b/thys/Spec.thy Thu Feb 25 22:46:58 2021 +0000 @@ -3,10 +3,7 @@ imports RegLangs begin - - - -section {* "Plain" Values *} +section \"Plain" Values\ datatype val = Void @@ -17,7 +14,7 @@ | Stars "val list" -section {* The string behind a value *} +section \The string behind a value\ fun flat :: "val \ string" @@ -38,7 +35,7 @@ by (induct vs) (auto) -section {* Lexical Values *} +section \Lexical Values\ inductive Prf :: "val \ rexp \ bool" ("\ _ : _" [100, 100] 100) @@ -47,7 +44,7 @@ | "\ v1 : r1 \ \ Left v1 : ALT r1 r2" | "\ v2 : r2 \ \ Right v2 : ALT r1 r2" | "\ Void : ONE" -| "\ Char c : CHAR c" +| "\ Char c : CH c" | "\v \ set vs. \ v : r \ flat v \ [] \ \ Stars vs : STAR r" inductive_cases Prf_elims: @@ -55,7 +52,7 @@ "\ v : SEQ r1 r2" "\ v : ALT r1 r2" "\ v : ONE" - "\ v : CHAR c" + "\ v : CH c" "\ vs : STAR r" lemma Prf_Stars_appendE: @@ -118,11 +115,11 @@ -section {* Sets of Lexical Values *} +section \Sets of Lexical Values\ -text {* +text \ Shows that lexical values are finite for a given regex and string. -*} +\ definition LV :: "rexp \ string \ val set" @@ -131,7 +128,7 @@ lemma LV_simps: shows "LV ZERO s = {}" and "LV ONE s = (if s = [] then {Void} else {})" - and "LV (CHAR c) s = (if s = [c] then {Char c} else {})" + and "LV (CH c) s = (if s = [c] then {Char c} else {})" and "LV (ALT r1 r2) s = Left ` LV r1 s \ Right ` LV r2 s" unfolding LV_def by (auto intro: Prf.intros elim: Prf.cases) @@ -229,8 +226,8 @@ case (ONE s) show "finite (LV ONE s)" by (simp add: LV_simps) next - case (CHAR c s) - show "finite (LV (CHAR c) s)" by (simp add: LV_simps) + case (CH c s) + show "finite (LV (CH c) s)" by (simp add: LV_simps) next case (ALT r1 r2 s) then show "finite (LV (ALT r1 r2) s)" by (simp add: LV_simps) @@ -258,13 +255,13 @@ -section {* Our inductive POSIX Definition *} +section \Our inductive POSIX Definition\ inductive Posix :: "string \ rexp \ val \ bool" ("_ \ _ \ _" [100, 100, 100] 100) where Posix_ONE: "[] \ ONE \ Void" -| Posix_CHAR: "[c] \ (CHAR c) \ (Char c)" +| Posix_CH: "[c] \ (CH c) \ (Char c)" | Posix_ALT1: "s \ r1 \ v \ s \ (ALT r1 r2) \ (Left v)" | Posix_ALT2: "\s \ r2 \ v; s \ L(r1)\ \ s \ (ALT r1 r2) \ (Right v)" | Posix_SEQ: "\s1 \ r1 \ v1; s2 \ r2 \ v2; @@ -278,7 +275,7 @@ inductive_cases Posix_elims: "s \ ZERO \ v" "s \ ONE \ v" - "s \ CHAR c \ v" + "s \ CH c \ v" "s \ ALT r1 r2 \ v" "s \ SEQ r1 r2 \ v" "s \ STAR r \ v" @@ -287,13 +284,13 @@ assumes "s \ r \ v" shows "s \ L r" "flat v = s" using assms -by (induct s r v rule: Posix.induct) - (auto simp add: Sequ_def) + by(induct s r v rule: Posix.induct) + (auto simp add: Sequ_def) -text {* +text \ For a give value and string, our Posix definition determines a unique value. -*} +\ lemma Posix_determ: assumes "s \ r \ v1" "s \ r \ v2" @@ -304,8 +301,8 @@ have "[] \ ONE \ v2" by fact then show "Void = v2" by cases auto next - case (Posix_CHAR c v2) - have "[c] \ CHAR c \ v2" by fact + case (Posix_CH c v2) + have "[c] \ CH c \ v2" by fact then show "Char c = v2" by cases auto next case (Posix_ALT1 s r1 v r2 v2) @@ -362,9 +359,9 @@ qed -text {* +text \ Our POSIX values are lexical values. -*} +\ lemma Posix_LV: assumes "s \ r \ v"