# HG changeset patch # User cu # Date 1507410976 -3600 # Node ID a3134f7de0654036b1387c8021929a10f44bc1c9 # Parent deea42c83c9e541ca8649ea0302299a1fa8976ea updated diff -r deea42c83c9e -r a3134f7de065 thys/LexerExt.thy --- a/thys/LexerExt.thy Thu Oct 05 12:45:13 2017 +0100 +++ b/thys/LexerExt.thy Sat Oct 07 22:16:16 2017 +0100 @@ -1,613 +1,23 @@ theory LexerExt - imports Main + imports SpecExt begin -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 *} - -lemma Sequ_empty [simp]: - shows "A ;; {[]} = A" - and "{[]} ;; A = A" -by (simp_all add: Sequ_def) - -lemma Sequ_null [simp]: - shows "A ;; {} = {}" - and "{} ;; A = {}" -by (simp_all add: Sequ_def) - -lemma Sequ_assoc: - shows "A ;; (B ;; C) = (A ;; B) ;; C" -apply(auto simp add: Sequ_def) -apply(metis append_assoc) -apply(metis) -done - -lemma Sequ_UNION: - shows "(\x\A. B ;; C x) = B ;; (\x\A. C x)" -by (auto simp add: Sequ_def) - - -section {* Semantic Derivative (Left Quotient) of Languages *} - -definition - Der :: "char \ string set \ string set" -where - "Der c A \ {s. c # s \ A}" - -definition - Ders :: "string \ string set \ string set" -where - "Ders s A \ {s'. s @ s' \ A}" - -lemma Der_null [simp]: - shows "Der c {} = {}" -unfolding Der_def -by auto - -lemma Der_empty [simp]: - shows "Der c {[]} = {}" -unfolding Der_def -by auto - -lemma Der_char [simp]: - shows "Der c {[d]} = (if c = d then {[]} else {})" -unfolding Der_def -by auto - -lemma Der_Sequ [simp]: - shows "Der c (A ;; B) = (Der c A) ;; B \ (if [] \ A then Der c B else {})" -unfolding Der_def Sequ_def -by (auto simp add: Cons_eq_append_conv) - -lemma Der_union [simp]: - shows "Der c (A \ B) = Der c A \ Der c B" -unfolding Der_def -by auto - -lemma Der_UNION: - shows "Der c (\x\A. B x) = (\x\A. Der c (B x))" -by (auto simp add: Der_def) - - -section {* Power operation for Sets *} - -fun - Pow :: "string set \ nat \ string set" ("_ \ _" [101, 102] 101) -where - "A \ 0 = {[]}" -| "A \ (Suc n) = A ;; (A \ n)" - -lemma Pow_empty [simp]: - shows "[] \ A \ n \ (n = 0 \ [] \ A)" -by(induct n) (auto simp add: Sequ_def) - -lemma Der_Pow [simp]: - shows "Der c (A \ (Suc n)) = - (Der c A) ;; (A \ n) \ (if [] \ A then Der c (A \ n) else {})" -unfolding Der_def Sequ_def -by(auto simp add: Cons_eq_append_conv Sequ_def) - -lemma Der_Pow_subseteq: - assumes "[] \ A" - shows "Der c (A \ n) \ (Der c A) ;; (A \ n)" -using assms -apply(induct n) -apply(simp add: Sequ_def Der_def) -apply(simp) -apply(rule conjI) -apply (smt Sequ_def append_Nil2 mem_Collect_eq Sequ_assoc subsetI) -apply(subgoal_tac "((Der c A) ;; (A \ n)) \ ((Der c A) ;; (A ;; (A \ n)))") -apply(auto)[1] -by (smt Sequ_def append_Nil2 mem_Collect_eq Sequ_assoc subsetI) - -lemma test: - shows "(\x\Suc n. Der c (A \ x)) = (\x\n. Der c A ;; A \ x)" -apply(induct n) -apply(simp) -apply(auto)[1] -apply(case_tac xa) -apply(simp) -apply(simp) -apply(auto)[1] -apply(case_tac "[] \ A") -apply(simp) -apply(simp) -by (smt Der_Pow Der_Pow_subseteq UN_insert atMost_Suc sup.orderE sup_bot.right_neutral) - -lemma test2: - shows "(\x\(Suc ` A). Der c (B \ x)) = (\x\A. Der c B ;; B \ x)" -apply(auto)[1] -apply(case_tac "[] \ B") -apply(simp) -using Der_Pow_subseteq apply blast -apply(simp) -done - - -section {* Kleene Star for Languages *} - -definition - Star :: "string set \ string set" ("_\" [101] 102) where - "A\ = (\n. A \ n)" - -lemma Star_empty [intro]: - shows "[] \ A\" -unfolding Star_def -by auto - -lemma Star_step [intro]: - assumes "s1 \ A" "s2 \ A\" - shows "s1 @ s2 \ A\" -proof - - from assms obtain n where "s1 \A" "s2 \ A \ n" - unfolding Star_def by auto - then have "s1 @ s2 \ A ;; (A \ n)" - by (auto simp add: Sequ_def) - then have "s1 @ s2 \ A \ (Suc n)" - by simp - then show "s1 @ s2 \ A\" - unfolding Star_def - by (auto simp del: Pow.simps) -qed - -lemma star_cases: - shows "A\ = {[]} \ A ;; A\" -unfolding Star_def -apply(simp add: Sequ_def) -apply(auto) -apply(case_tac xa) -apply(auto simp add: Sequ_def) -apply(rule_tac x="Suc xa" in exI) -apply(auto simp add: Sequ_def) -done - -lemma Der_Star1: - shows "Der c (A ;; A\) = (Der c A) ;; A\" -proof - - have "(Der c A) ;; A\ = (Der c A) ;; (\n. A \ n)" - unfolding Star_def by simp - also have "... = (\n. Der c A ;; A \ n)" - unfolding Sequ_UNION by simp - also have "... = (\x\(Suc ` UNIV). Der c (A \ x))" - unfolding test2 by simp - also have "... = (\n. Der c (A \ (Suc n)))" - by (simp) - also have "... = Der c (\n. A \ (Suc n))" - unfolding Der_UNION by simp - also have "... = Der c (A ;; (\n. A \ n))" - by (simp only: Pow.simps Sequ_UNION) - finally show "Der c (A ;; A\) = (Der c A) ;; A\" - unfolding Star_def[symmetric] by blast -qed - -lemma Der_star [simp]: - shows "Der c (A\) = (Der c A) ;; A\" -proof - - have "Der c (A\) = Der c ({[]} \ A ;; A\)" - by (simp only: star_cases[symmetric]) - also have "... = Der c (A ;; A\)" - by (simp only: Der_union Der_empty) (simp) - also have "... = (Der c A) ;; A\" - using Der_Star1 by simp - finally show "Der c (A\) = (Der c A) ;; A\" . -qed - - - - -section {* Regular Expressions *} - -datatype rexp = - ZERO -| ONE -| CHAR char -| SEQ rexp rexp -| ALT rexp rexp -| STAR rexp -| UPNTIMES rexp nat -| NTIMES rexp nat -| FROMNTIMES rexp nat -| NMTIMES rexp nat nat -| PLUS rexp - -section {* Semantics of Regular Expressions *} - -fun - L :: "rexp \ string set" -where - "L (ZERO) = {}" -| "L (ONE) = {[]}" -| "L (CHAR c) = {[c]}" -| "L (SEQ r1 r2) = (L r1) ;; (L r2)" -| "L (ALT r1 r2) = (L r1) \ (L r2)" -| "L (STAR r) = (L r)\" -| "L (UPNTIMES r n) = (\i\ {..n} . (L r) \ i)" -| "L (NTIMES r n) = (L r) \ n" -| "L (FROMNTIMES r n) = (\i\ {n..} . (L r) \ i)" -| "L (NMTIMES r n m) = (\i\{n..m} . (L r) \ i)" -| "L (PLUS r) = (\i\ {1..} . (L r) \ i)" - -section {* Nullable, Derivatives *} - -fun - nullable :: "rexp \ bool" -where - "nullable (ZERO) = False" -| "nullable (ONE) = True" -| "nullable (CHAR c) = False" -| "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" -| "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" -| "nullable (STAR r) = True" -| "nullable (UPNTIMES r n) = True" -| "nullable (NTIMES r n) = (if n = 0 then True else nullable r)" -| "nullable (FROMNTIMES r n) = (if n = 0 then True else nullable r)" -| "nullable (NMTIMES r n m) = (if m < n then False else (if n = 0 then True else nullable r))" -| "nullable (PLUS r) = (nullable r)" - -fun - der :: "char \ rexp \ rexp" -where - "der c (ZERO) = ZERO" -| "der c (ONE) = ZERO" -| "der c (CHAR 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 - then ALT (SEQ (der c r1) r2) (der c r2) - else SEQ (der c r1) r2)" -| "der c (STAR r) = SEQ (der c r) (STAR r)" -| "der c (UPNTIMES r 0) = ZERO" -| "der c (UPNTIMES r (Suc n)) = SEQ (der c r) (UPNTIMES r n)" -| "der c (NTIMES r 0) = ZERO" -| "der c (NTIMES r (Suc n)) = SEQ (der c r) (NTIMES r n)" -| "der c (FROMNTIMES r 0) = SEQ (der c r) (FROMNTIMES r 0)" -| "der c (FROMNTIMES r (Suc n)) = SEQ (der c r) (FROMNTIMES r n)" -| "der c (NMTIMES r n m) = - (if m < n then ZERO - else (if n = 0 then (if m = 0 then ZERO else - SEQ (der c r) (UPNTIMES r (m - 1))) else - SEQ (der c r) (NMTIMES r (n - 1) (m - 1))))" -| "der c (PLUS r) = SEQ (der c r) (STAR r)" - -fun - ders :: "string \ rexp \ rexp" -where - "ders [] r = r" -| "ders (c # s) r = ders s (der c r)" - - -lemma nullable_correctness: - shows "nullable r \ [] \ (L r)" -apply(induct r) -apply(auto simp add: Sequ_def) -done - -lemma Der_Pow_notin: - assumes "[] \ A" - shows "Der c (A \ (Suc n)) = (Der c A) ;; (A \ n)" -using assms -by(simp) - -lemma der_correctness: - shows "L (der c r) = Der c (L r)" -apply(induct c r rule: der.induct) -apply(simp_all add: nullable_correctness)[7] -apply(simp only: der.simps L.simps) -apply(simp only: Der_UNION) -apply(simp only: Sequ_UNION[symmetric]) -apply(simp add: test) -apply(simp) -(* NTIMES *) -apply(simp only: L.simps der.simps) -apply(simp) -apply(rule impI) -apply (simp add: Der_Pow_subseteq sup_absorb1) -(* FROMNTIMES *) -apply(simp only: der.simps) -apply(simp only: L.simps) -apply(simp) -using Der_star Star_def apply auto[1] -apply(simp only: der.simps) -apply(simp only: L.simps) -apply(simp add: Der_UNION) -apply(subst Sequ_UNION[symmetric]) -apply(subst test2[symmetric]) -apply(subgoal_tac "(Suc ` {n..}) = {Suc n..}") -apply simp -apply(auto simp add: image_def)[1] -using Suc_le_D apply blast -(* NMTIMES *) -apply(case_tac "n \ m") -prefer 2 -apply(simp only: der.simps if_True) -apply(simp only: L.simps) -apply(simp) -apply(auto) -apply(subst (asm) Sequ_UNION[symmetric]) -apply(subst (asm) test[symmetric]) -apply(auto simp add: Der_UNION)[1] -apply(auto simp add: Der_UNION)[1] -apply(subst Sequ_UNION[symmetric]) -apply(subst test[symmetric]) -apply(auto)[1] -apply(subst (asm) Sequ_UNION[symmetric]) -apply(subst (asm) test2[symmetric]) -apply(auto simp add: Der_UNION)[1] -apply(subst Sequ_UNION[symmetric]) -apply(subst test2[symmetric]) -apply(auto simp add: Der_UNION)[1] -(* PLUS *) -apply(auto simp add: Sequ_def Star_def)[1] -apply(simp add: Der_UNION) -apply(rule_tac x="Suc xa" in bexI) -apply(auto simp add: Sequ_def Der_def)[2] -apply (metis append_Cons) -apply(simp add: Der_UNION) -apply(erule_tac bexE) -apply(case_tac xa) -apply(simp) -apply(simp) -apply(auto) -apply(auto simp add: Sequ_def Der_def)[1] -using Star_def apply auto[1] -apply(case_tac "[] \ L r") -apply(auto) -using Der_UNION Der_star Star_def by fastforce - - -lemma ders_correctness: - shows "L (ders s r) = Ders s (L r)" -apply(induct s arbitrary: r) -apply(simp_all add: Ders_def der_correctness Der_def) -done - -lemma ders_ZERO: - shows "ders s (ZERO) = ZERO" -apply(induct s) -apply(simp_all) -done - -lemma ders_ONE: - shows "ders s (ONE) = (if s = [] then ONE else ZERO)" -apply(induct s) -apply(simp_all add: ders_ZERO) -done - -lemma ders_CHAR: - shows "ders s (CHAR c) = (if s = [c] then ONE else - (if s = [] then (CHAR c) else ZERO))" -apply(induct s) -apply(simp_all add: ders_ZERO ders_ONE) -done - -lemma ders_ALT: - shows "ders s (ALT r1 r2) = ALT (ders s r1) (ders s r2)" -apply(induct s arbitrary: r1 r2) -apply(simp_all) -done - -section {* Values *} - -datatype val = - Void -| Char char -| Seq val val -| Right val -| Left val -| Stars "val list" - - -section {* The string behind a value *} - -fun - flat :: "val \ string" -where - "flat (Void) = []" -| "flat (Char c) = [c]" -| "flat (Left v) = flat v" -| "flat (Right v) = flat v" -| "flat (Seq v1 v2) = (flat v1) @ (flat v2)" -| "flat (Stars []) = []" -| "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" - -lemma flat_Stars [simp]: - "flat (Stars vs) = concat (map flat vs)" -by (induct vs) (auto) - - -section {* Relation between values and regular expressions *} - -inductive - Prf :: "val \ rexp \ bool" ("\ _ : _" [100, 100] 100) -where - "\\ v1 : r1; \ v2 : r2\ \ \ Seq v1 v2 : SEQ r1 r2" -| "\ v1 : r1 \ \ Left v1 : ALT r1 r2" -| "\ v2 : r2 \ \ Right v2 : ALT r1 r2" -| "\ Void : ONE" -| "\ Char c : CHAR c" -| "\\v \ set vs. \ v : r\ \ \ Stars vs : STAR r" -| "\\v \ set vs. \ v : r; length vs \ n\ \ \ Stars vs : UPNTIMES r n" -| "\\v \ set vs. \ v : r; length vs = n\ \ \ Stars vs : NTIMES r n" -| "\\v \ set vs. \ v : r; length vs \ n\ \ \ Stars vs : FROMNTIMES r n" -| "\\v \ set vs. \ v : r; length vs \ n; length vs \ m\ \ \ Stars vs : NMTIMES r n m" -| "\\v \ set vs. \ v : r; length vs \ 1\ \ \ Stars vs : PLUS r" - - -inductive_cases Prf_elims: - "\ v : ZERO" - "\ v : SEQ r1 r2" - "\ v : ALT r1 r2" - "\ v : ONE" - "\ v : CHAR c" -(* "\ vs : STAR r"*) - -lemma Prf_STAR: - assumes "\v\set vs. \ v : r \ flat v \ L r" - shows "concat (map flat vs) \ L r\" -using assms -apply(induct vs) -apply(auto) -done - -lemma Prf_Pow: - assumes "\v\set vs. \ v : r \ flat v \ L r" - shows "concat (map flat vs) \ L r \ length vs" -using assms -apply(induct vs) -apply(auto simp add: Sequ_def) -done - -lemma Prf_flat_L: - assumes "\ v : r" shows "flat v \ L r" -using assms -apply(induct v r rule: Prf.induct) -apply(auto simp add: Sequ_def) -apply(rule Prf_STAR) -apply(simp) -apply(rule_tac x="length vs" in bexI) -apply(rule Prf_Pow) -apply(simp) -apply(simp) -apply(rule Prf_Pow) -apply(simp) -apply(rule_tac x="length vs" in bexI) -apply(rule Prf_Pow) -apply(simp) -apply(simp) -apply(rule_tac x="length vs" in bexI) -apply(rule Prf_Pow) -apply(simp) -apply(simp) -using Prf_Pow by blast - -lemma Star_Pow: - assumes "s \ A \ n" - shows "\ss. concat ss = s \ (\s \ set ss. s \ A) \ (length ss = n)" -using assms -apply(induct n arbitrary: s) -apply(auto simp add: Sequ_def) -apply(drule_tac x="s2" in meta_spec) -apply(auto) -apply(rule_tac x="s1#ss" in exI) -apply(simp) -done - -lemma Star_string: - assumes "s \ A\" - shows "\ss. concat ss = s \ (\s \ set ss. s \ A)" -using assms -apply(auto simp add: Star_def) -using Star_Pow by blast - - -lemma Star_val: - assumes "\s\set ss. \v. s = flat v \ \ v : r" - shows "\vs. concat (map flat vs) = concat ss \ (\v\set vs. \ v : r)" -using assms -apply(induct ss) -apply(auto) -apply (metis empty_iff list.set(1)) -by (metis concat.simps(2) list.simps(9) set_ConsD) - -lemma Star_val_length: - assumes "\s\set ss. \v. s = flat v \ \ v : r" - shows "\vs. concat (map flat vs) = concat ss \ (\v\set vs. \ v : r) \ (length vs) = (length ss)" -using assms -apply(induct ss) -apply(auto) -by (metis List.bind_def bind_simps(2) length_Suc_conv set_ConsD) - - - - - -lemma L_flat_Prf2: - assumes "s \ L r" shows "\v. \ v : r \ flat v = s" -using assms -apply(induct r arbitrary: s) -apply(auto simp add: Sequ_def intro: Prf.intros) -using Prf.intros(1) flat.simps(5) apply blast -using Prf.intros(2) flat.simps(3) apply blast -using Prf.intros(3) flat.simps(4) apply blast -apply(subgoal_tac "\vs::val list. concat (map flat vs) = s \ (\v \ set vs. \ v : r)") -apply(auto)[1] -apply(rule_tac x="Stars vs" in exI) -apply(simp) -apply(rule Prf.intros) -apply(simp) -using Star_string Star_val apply force -apply(subgoal_tac "\vs::val list. concat (map flat vs) = s \ (\v \ set vs. \ v : r) \ (length vs = x)") -apply(auto)[1] -apply(rule_tac x="Stars vs" in exI) -apply(simp) -apply(rule Prf.intros) -apply(simp) -apply(simp) -using Star_Pow Star_val_length apply blast -apply(subgoal_tac "\vs::val list. concat (map flat vs) = s \ (\v \ set vs. \ v : r) \ (length vs = x2)") -apply(auto)[1] -apply(rule_tac x="Stars vs" in exI) -apply(simp) -apply(rule Prf.intros) -apply(simp) -apply(simp) -using Star_Pow Star_val_length apply blast -apply(subgoal_tac "\vs::val list. concat (map flat vs) = s \ (\v \ set vs. \ v : r) \ (length vs = x)") -apply(auto)[1] -apply(rule_tac x="Stars vs" in exI) -apply(simp) -apply(rule Prf.intros) -apply(simp) -apply(simp) -using Star_Pow Star_val_length apply blast -apply(subgoal_tac "\vs::val list. concat (map flat vs) = s \ (\v \ set vs. \ v : r) \ (length vs = x)") -apply(auto)[1] -apply(rule_tac x="Stars vs" in exI) -apply(simp) -apply(rule Prf.intros) -apply(simp) -apply(simp) -apply(simp) -using Star_Pow Star_val_length apply blast -apply(subgoal_tac "\vs::val list. concat (map flat vs) = s \ (\v \ set vs. \ v : r) \ (length vs = x)") -apply(auto)[1] -apply(rule_tac x="Stars vs" in exI) -apply(simp) -apply(rule Prf.intros) -apply(simp) -apply(simp) -using Star_Pow Star_val_length apply blast -done - -lemma L_flat_Prf: - "L(r) = {flat v | v. \ v : r}" -using Prf_flat_L L_flat_Prf2 by blast - - -section {* Sulzmann and Lu functions *} +section {* The Lexer Functions by Sulzmann and Lu *} fun mkeps :: "rexp \ val" where "mkeps(ONE) = Void" | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)" -| "mkeps(ALT r1 r2) = (if nullable r1 then Left (mkeps r1) else Right (mkeps r2))" +| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))" | "mkeps(STAR r) = Stars []" | "mkeps(UPNTIMES r n) = Stars []" | "mkeps(NTIMES r n) = Stars (replicate n (mkeps r))" | "mkeps(FROMNTIMES r n) = Stars (replicate n (mkeps r))" | "mkeps(NMTIMES r n m) = Stars (replicate n (mkeps r))" -| "mkeps(PLUS r) = Stars ([mkeps r])" - - + fun injval :: "rexp \ char \ val \ val" where "injval (CHAR d) c Void = Char d" @@ -617,364 +27,157 @@ | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2" | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)" | "injval (STAR r) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" -| "injval (UPNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" -| "injval (NTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" -| "injval (FROMNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" -| "injval (NMTIMES r n m) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" -| "injval (PLUS r) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" - -section {* Mkeps, injval *} +| "injval (NTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" +| "injval (FROMNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" +| "injval (UPNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" +| "injval (NMTIMES r n m) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" + +fun + lexer :: "rexp \ string \ val option" +where + "lexer r [] = (if nullable r then Some(mkeps r) else None)" +| "lexer r (c#s) = (case (lexer (der c r) s) of + None \ None + | Some(v) \ Some(injval r c v))" -lemma mkeps_nullable: - assumes "nullable(r)" - shows "\ mkeps r : r" -using assms -apply(induct r rule: mkeps.induct) -apply(auto intro: Prf.intros) -by (metis Prf.intros(10) in_set_replicate length_replicate not_le order_refl) + + +section {* Mkeps, Injval Properties *} lemma mkeps_flat: assumes "nullable(r)" shows "flat (mkeps r) = []" using assms -apply (induct rule: nullable.induct) -apply(auto) -by meson - - -lemma Prf_injval: - assumes "\ v : der c r" - shows "\ (injval r c v) : r" + apply(induct rule: nullable.induct) + apply(auto) + by presburger + + +lemma mkeps_nullable: + assumes "nullable(r)" + shows "\ mkeps r : r" using assms -apply(induct r arbitrary: c v rule: rexp.induct) -apply(auto intro!: Prf.intros mkeps_nullable elim!: Prf_elims split: if_splits) -(* STAR *) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(auto) -using Prf.intros(6) apply auto[1] -(* UPNTIMES *) -apply(case_tac x2) -apply(simp) -using Prf_elims(1) apply auto[1] -apply(simp) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -using Prf.intros(7) apply auto[1] -(* NTIMES *) -apply(case_tac x2) -apply(simp) -using Prf_elims(1) apply auto[1] -apply(simp) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -using Prf.intros(8) apply auto[1] -(* FROMNTIMES *) -apply(case_tac x2) -apply(simp) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -using Prf.intros(9) apply auto[1] -apply(rotate_tac 1) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -using Prf.intros(9) apply auto -(* NMTIMES *) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -apply(rule Prf.intros) -apply(auto)[2] -apply simp -apply(rotate_tac 4) -apply(erule Prf.cases) -apply(simp_all) -apply(clarify) -apply(rule Prf.intros) -apply(auto)[2] -apply simp -(* PLUS *) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(auto) -using Prf.intros apply auto[1] +apply(induct rule: nullable.induct) + apply(auto intro: Prf.intros split: if_splits) + using Prf.intros(8) apply force + apply(subst append.simps(1)[symmetric]) + apply(rule Prf.intros) + apply(simp) + apply(simp) + apply (simp add: mkeps_flat) + apply(simp) + using Prf.intros(9) apply force + apply(subst append.simps(1)[symmetric]) + apply(rule Prf.intros) + apply(simp) + apply(simp) + apply (simp add: mkeps_flat) + apply(simp) + using Prf.intros(11) apply force + apply(subst append.simps(1)[symmetric]) + apply(rule Prf.intros) + apply(simp) + apply(simp) + apply (simp add: mkeps_flat) + apply(simp) + apply(simp) done - + lemma Prf_injval_flat: - assumes "\ v : der c r" + assumes "\ v : der c r" shows "flat (injval r c v) = c # (flat v)" using assms apply(induct arbitrary: v rule: der.induct) -apply(auto elim!: Prf_elims split: if_splits) -apply(metis mkeps_flat) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 4) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 2) -apply(erule Prf.cases) -apply(simp_all) +apply(auto elim!: Prf_elims intro: mkeps_flat split: if_splits) +done + +lemma Prf_injval: + assumes "\ v : der c r" + shows "\ (injval r c v) : r" +using assms +apply(induct r arbitrary: c v rule: rexp.induct) +apply(auto intro!: Prf.intros mkeps_nullable elim!: Prf_elims split: if_splits)[6] + apply(simp add: Prf_injval_flat) + apply(simp) + apply(case_tac x2) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + using Prf.intros(7) Prf_injval_flat apply auto[1] + apply(simp) + apply(case_tac x2) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(subst append.simps(2)[symmetric]) + apply(rule Prf.intros) + apply(simp add: Prf_injval_flat) + apply(simp) + apply(simp) + prefer 2 + apply(simp) + apply(case_tac "x3a < x2") + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(case_tac x2) + apply(simp) + apply(case_tac x3a) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + using Prf.intros(12) Prf_injval_flat apply auto[1] + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(subst append.simps(2)[symmetric]) + apply(rule Prf.intros) + apply(simp add: Prf_injval_flat) + apply(simp) + apply(simp) + apply(simp) + apply(simp) + using Prf.intros(12) Prf_injval_flat apply auto[1] + apply(case_tac x2) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp_all) + apply (simp add: Prf.intros(10) Prf_injval_flat) + using Prf.intros(10) Prf_injval_flat apply auto[1] + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp_all) + apply(subst append.simps(2)[symmetric]) + apply(rule Prf.intros) + apply(simp add: Prf_injval_flat) + apply(simp) + apply(simp) + using Prf.intros(10) Prf_injval_flat apply auto done -section {* Our Alternative 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_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; - \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r1 \ s\<^sub>4 \ L r2)\ \ - (s1 @ s2) \ (SEQ r1 r2) \ (Seq v1 v2)" -| Posix_STAR1: "\s1 \ r \ v; s2 \ STAR r \ Stars vs; flat v \ []; - \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (STAR r))\ - \ (s1 @ s2) \ STAR r \ Stars (v # vs)" -| Posix_STAR2: "[] \ STAR r \ Stars []" -| Posix_UPNTIMES1: "\s1 \ r \ v; s2 \ UPNTIMES r n \ Stars vs; flat v \ []; - \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (UPNTIMES r n))\ - \ (s1 @ s2) \ UPNTIMES r (Suc n) \ Stars (v # vs)" -| Posix_UPNTIMES2: "[] \ UPNTIMES r n \ Stars []" -| Posix_NTIMES1: "\s1 \ r \ v; s2 \ NTIMES r n \ Stars vs; flat v = [] \ flat (Stars vs) = []; - \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (NTIMES r n))\ - \ (s1 @ s2) \ NTIMES r (Suc n) \ Stars (v # vs)" -| Posix_NTIMES2: "[] \ NTIMES r 0 \ Stars []" -| Posix_FROMNTIMES1: "\s1 \ r \ v; s2 \ FROMNTIMES r n \ Stars vs; flat v = [] \ flat (Stars vs) = []; - \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (FROMNTIMES r n))\ - \ (s1 @ s2) \ FROMNTIMES r (Suc n) \ Stars (v # vs)" -| Posix_FROMNTIMES2: "s \ STAR r \ Stars vs \ s \ FROMNTIMES r 0 \ Stars vs" -| Posix_NMTIMES1: "\s1 \ r \ v; s2 \ NMTIMES r n m \ Stars vs; flat v = [] \ flat (Stars vs) = []; - \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (NMTIMES r n m))\ - \ (s1 @ s2) \ NMTIMES r (Suc n) (Suc m) \ Stars (v # vs)" -| Posix_NMTIMES2: "s \ UPNTIMES r m \ Stars vs \ s \ NMTIMES r 0 m \ Stars vs" -| Posix_PLUS1: "\s1 \ r \ v; s2 \ STAR r \ Stars vs; flat v = [] \ flat (Stars vs) = []; - \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (STAR r))\ - \ (s1 @ s2) \ PLUS r \ Stars (v # vs)" - -inductive_cases Posix_elims: - "s \ ZERO \ v" - "s \ ONE \ v" - "s \ CHAR c \ v" - "s \ ALT r1 r2 \ v" - "s \ SEQ r1 r2 \ v" - "s \ STAR r \ v" - -lemma Posix1: - assumes "s \ r \ v" - shows "s \ L r" "flat v = s" -using assms -apply (induct s r v rule: Posix.induct) -apply(auto simp add: Sequ_def) -apply(rule_tac x="Suc x" in bexI) -apply(auto simp add: Sequ_def) -apply(rule_tac x="Suc x" in bexI) -using Sequ_def apply auto[2] -apply(simp add: Star_def) -apply(rule_tac x="Suc x" in bexI) -apply(auto simp add: Sequ_def) -apply(simp add: Star_def) -apply(clarify) -apply(rule_tac x="Suc x" in bexI) -apply(auto simp add: Sequ_def) -done - - -lemma - "([] @ []) \ PLUS (ONE) \ Stars ([Void])" -apply(rule Posix_PLUS1) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(simp) -apply(simp) -done - -lemma - assumes "s \ r \ v" "flat v \ []" "\s' \ L r. length s' < length s" - shows "([] @ (s @ [])) \ PLUS (ALT ONE r) \ Stars ([Left Void, Right v])" -using assms -oops - -lemma - "([] @ ([] @ [])) \ FROMNTIMES (ONE) (Suc (Suc 0)) \ Stars ([Void, Void])" -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(auto) -done - - -lemma - assumes "s \ r \ v" "flat v \ []" - "s \ PLUS (ALT ONE r) \ Stars ([Left(Void), Right(v)])" - shows "False" -using assms -apply(rotate_tac 2) -apply(erule_tac Posix.cases) -apply(simp_all) -apply(auto) -oops - - - - - -lemma Posix1a: - assumes "s \ r \ v" - shows "\ v : r" -using assms -apply(induct s r v rule: Posix.induct) -apply(auto intro: Prf.intros) -apply(rule Prf.intros) -apply(auto)[1] -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rule Prf.intros) -apply(auto)[1] -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rule Prf.intros) -apply(auto)[1] -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rule Prf.intros) -apply(auto)[1] -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rule Prf.intros) -apply(auto)[2] -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rule Prf.intros) -apply(auto)[3] -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rule Prf.intros) -apply(auto)[3] -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(erule Prf.cases) -apply(simp_all) -apply(rotate_tac 3) -apply(erule Prf.cases) -apply(simp_all) -apply(rule Prf.intros) -apply(auto) -done - - -lemma Posix_NTIMES_mkeps: - assumes "[] \ r \ mkeps r" - shows "[] \ NTIMES r n \ Stars (replicate n (mkeps r))" -apply(induct n) -apply(simp) -apply (rule Posix_NTIMES2) -apply(simp) -apply(subst append_Nil[symmetric]) -apply (rule Posix_NTIMES1) -apply(auto) -apply(rule assms) -done - -lemma Posix_FROMNTIMES_mkeps: - assumes "[] \ r \ mkeps r" - shows "[] \ FROMNTIMES r n \ Stars (replicate n (mkeps r))" -apply(induct n) -apply(simp) -apply (rule Posix_FROMNTIMES2) -apply (rule Posix.intros) -apply(simp) -apply(subst append_Nil[symmetric]) -apply (rule Posix_FROMNTIMES1) -apply(auto) -apply(rule assms) -done - -lemma Posix_NMTIMES_mkeps: - assumes "[] \ r \ mkeps r" "n \ m" - shows "[] \ NMTIMES r n m \ Stars (replicate n (mkeps r))" -using assms(2) -apply(induct n arbitrary: m) -apply(simp) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(case_tac m) -apply(simp) -apply(simp) -apply(subst append_Nil[symmetric]) -apply(rule Posix.intros) -apply(auto) -apply(rule assms) -done - - +text {* + Mkeps and injval produce, or preserve, Posix values. +*} lemma Posix_mkeps: assumes "nullable r" @@ -984,261 +187,21 @@ apply(auto intro: Posix.intros simp add: nullable_correctness Sequ_def) apply(subst append.simps(1)[symmetric]) apply(rule Posix.intros) -apply(auto) -apply(case_tac n) -apply(simp) -apply (simp add: Posix_NTIMES2) -apply(simp) -apply(subst append.simps(1)[symmetric]) -apply(rule Posix.intros) -apply(auto) -apply(rule Posix_NTIMES_mkeps) -apply(simp) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(case_tac n) -apply(simp) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(simp) -apply(subst append.simps(1)[symmetric]) -apply(rule Posix.intros) -apply(auto) -apply(rule Posix_FROMNTIMES_mkeps) -apply(simp) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(case_tac n) -apply(simp) -apply(rule Posix.intros) -apply(rule Posix.intros) -apply(simp) -apply(case_tac m) -apply(simp) -apply(simp) -apply(subst append_Nil[symmetric]) -apply(rule Posix.intros) -apply(auto) -apply(rule Posix_NMTIMES_mkeps) -apply(auto) -apply(subst append_Nil[symmetric]) -apply(rule Posix.intros) -apply(auto) -apply(rule Posix.intros) -done - - -lemma Posix_determ: - assumes "s \ r \ v1" "s \ r \ v2" - shows "v1 = v2" -using assms -proof (induct s r v1 arbitrary: v2 rule: Posix.induct) - case (Posix_ONE v2) - 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 - then show "Char c = v2" by cases auto -next - case (Posix_ALT1 s r1 v r2 v2) - have "s \ ALT r1 r2 \ v2" by fact - moreover - have "s \ r1 \ v" by fact - then have "s \ L r1" by (simp add: Posix1) - ultimately obtain v' where eq: "v2 = Left v'" "s \ r1 \ v'" by cases auto - moreover - have IH: "\v2. s \ r1 \ v2 \ v = v2" by fact - ultimately have "v = v'" by simp - then show "Left v = v2" using eq by simp -next - case (Posix_ALT2 s r2 v r1 v2) - have "s \ ALT r1 r2 \ v2" by fact - moreover - have "s \ L r1" by fact - ultimately obtain v' where eq: "v2 = Right v'" "s \ r2 \ v'" - by cases (auto simp add: Posix1) - moreover - have IH: "\v2. s \ r2 \ v2 \ v = v2" by fact - ultimately have "v = v'" by simp - then show "Right v = v2" using eq by simp -next - case (Posix_SEQ s1 r1 v1 s2 r2 v2 v') - have "(s1 @ s2) \ SEQ r1 r2 \ v'" - "s1 \ r1 \ v1" "s2 \ r2 \ v2" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" by fact+ - then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \ r1 \ v1'" "s2 \ r2 \ v2'" - apply(cases) apply (auto simp add: append_eq_append_conv2) - using Posix1(1) by fastforce+ - moreover - have IHs: "\v1'. s1 \ r1 \ v1' \ v1 = v1'" - "\v2'. s2 \ r2 \ v2' \ v2 = v2'" by fact+ - ultimately show "Seq v1 v2 = v'" by simp -next - case (Posix_STAR1 s1 r v s2 vs v2) - have "(s1 @ s2) \ STAR r \ v2" - "s1 \ r \ v" "s2 \ STAR r \ Stars vs" "flat v \ []" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (STAR r))" by fact+ - then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (STAR r) \ (Stars vs')" - apply(cases) apply (auto simp add: append_eq_append_conv2) - using Posix1(1) apply fastforce - apply (metis Posix1(1) Posix_STAR1.hyps(6) append_Nil append_Nil2) - using Posix1(2) by blast - moreover - have IHs: "\v2. s1 \ r \ v2 \ v = v2" - "\v2. s2 \ STAR r \ v2 \ Stars vs = v2" by fact+ - ultimately show "Stars (v # vs) = v2" by auto -next - case (Posix_STAR2 r v2) - have "[] \ STAR r \ v2" by fact - then show "Stars [] = v2" by cases (auto simp add: Posix1) -next - case (Posix_UPNTIMES1 s1 r v s2 n vs v2) - have "(s1 @ s2) \ UPNTIMES r (Suc n) \ v2" - "s1 \ r \ v" "s2 \ (UPNTIMES r n) \ Stars vs" "flat v \ []" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (UPNTIMES r n))" by fact+ - then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (UPNTIMES r n) \ (Stars vs')" - apply(cases) apply (auto simp add: append_eq_append_conv2) - using Posix1(1) apply fastforce - apply (metis Posix1(1) Posix_UPNTIMES1.hyps(6) append_Nil append_Nil2) - using Posix1(2) by blast - moreover - have IHs: "\v2. s1 \ r \ v2 \ v = v2" - "\v2. s2 \ UPNTIMES r n \ v2 \ Stars vs = v2" by fact+ - ultimately show "Stars (v # vs) = v2" by auto -next - case (Posix_UPNTIMES2 r n v2) - have "[] \ UPNTIMES r n \ v2" by fact - then show "Stars [] = v2" by cases (auto simp add: Posix1) -next - case (Posix_NTIMES2 r v2) - have "[] \ NTIMES r 0 \ v2" by fact - then show "Stars [] = v2" by cases (auto simp add: Posix1) -next - case (Posix_NTIMES1 s1 r v s2 n vs v2) - have "(s1 @ s2) \ NTIMES r (Suc n) \ v2" "flat v = [] \ flat (Stars vs) = []" - "s1 \ r \ v" "s2 \ (NTIMES r n) \ Stars vs" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (NTIMES r n))" by fact+ - then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (NTIMES r n) \ (Stars vs')" - apply(cases) apply (auto simp add: append_eq_append_conv2) - using Posix1(1) apply fastforce - by (metis Posix1(1) Posix_NTIMES1.hyps(6) append_Nil append_Nil2) - moreover - have IHs: "\v2. s1 \ r \ v2 \ v = v2" - "\v2. s2 \ NTIMES r n \ v2 \ Stars vs = v2" by fact+ - ultimately show "Stars (v # vs) = v2" by auto -next - case (Posix_FROMNTIMES2 s r v1 v2) - have "s \ FROMNTIMES r 0 \ v2" "s \ STAR r \ Stars v1" - "\v3. s \ STAR r \ v3 \ Stars v1 = v3" by fact+ - then show ?case - apply(cases) + apply(auto) + done + +lemma test: + assumes "s \ der c (FROMNTIMES r 0) \ v" + shows "XXX" +using assms + apply(simp) + apply(erule Posix_elims) + apply(drule FROMNTIMES_0) apply(auto) - done -next - case (Posix_FROMNTIMES1 s1 r v s2 n vs v2) - have "(s1 @ s2) \ FROMNTIMES r (Suc n) \ v2" - "s1 \ r \ v" "s2 \ (FROMNTIMES r n) \ Stars vs" "flat v = [] \ flat (Stars vs) = []" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (FROMNTIMES r n))" by fact+ - then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (FROMNTIMES r n) \ (Stars vs')" - apply(cases) apply (auto simp add: append_eq_append_conv2) - using Posix1(1) apply fastforce - by (metis Posix1(1) Posix_FROMNTIMES1.hyps(6) append_Nil append_Nil2) - moreover - have IHs: "\v2. s1 \ r \ v2 \ v = v2" - "\v2. s2 \ FROMNTIMES r n \ v2 \ Stars vs = v2" by fact+ - ultimately show "Stars (v # vs) = v2" by auto -next - case (Posix_NMTIMES2 s r m v1 v2) - have "s \ NMTIMES r 0 m \ v2" "s \ UPNTIMES r m \ Stars v1" - "\v3. s \ UPNTIMES r m \ v3 \ Stars v1 = v3" by fact+ - then show ?case - apply(cases) - apply(auto) - done -next - case (Posix_NMTIMES1 s1 r v s2 n m vs v2) - have "(s1 @ s2) \ NMTIMES r (Suc n) (Suc m) \ v2" - "s1 \ r \ v" "s2 \ (NMTIMES r n m) \ Stars vs" "flat v = [] \ flat (Stars vs) = []" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (NMTIMES r n m))" by fact+ - then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (NMTIMES r n m) \ (Stars vs')" - apply(cases) apply (auto simp add: append_eq_append_conv2) - using Posix1(1) apply fastforce - by (metis Posix1(1) Posix_NMTIMES1.hyps(6) self_append_conv self_append_conv2) - moreover - have IHs: "\v2. s1 \ r \ v2 \ v = v2" - "\v2. s2 \ NMTIMES r n m \ v2 \ Stars vs = v2" by fact+ - ultimately show "Stars (v # vs) = v2" by auto -next - case (Posix_PLUS1 s1 r v s2 vs v2) - have "(s1 @ s2) \ PLUS r \ v2" - "s1 \ r \ v" "s2 \ (STAR r) \ Stars vs" (*"flat v = [] \ flat (Stars vs) = []"*) - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (STAR r))" by fact+ - then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (STAR r) \ (Stars vs')" - apply(cases) apply (auto simp add: append_eq_append_conv2) - using Posix1(1) apply fastforce - by (metis Posix1(1) Posix_PLUS1.hyps(6) append_self_conv append_self_conv2) - moreover - have IHs: "\v2. s1 \ r \ v2 \ v = v2" - "\v2. s2 \ STAR r \ v2 \ Stars vs = v2" by fact+ - ultimately show "Stars (v # vs) = v2" by auto -qed - -lemma NTIMES_Stars: - assumes "\ v : NTIMES r n" - shows "\vs. v = Stars vs \ (\v \ set vs. \ v : r) \ length vs = n" -using assms -apply(induct n arbitrary: v) -apply(erule Prf.cases) -apply(simp_all) -apply(erule Prf.cases) -apply(simp_all) -done - -lemma UPNTIMES_Stars: - assumes "\ v : UPNTIMES r n" - shows "\vs. v = Stars vs \ (\v \ set vs. \ v : r) \ length vs \ n" -using assms -apply(induct n arbitrary: v) -apply(erule Prf.cases) -apply(simp_all) -apply(erule Prf.cases) -apply(simp_all) -done - -lemma FROMNTIMES_Stars: - assumes "\ v : FROMNTIMES r n" - shows "\vs. v = Stars vs \ (\v \ set vs. \ v : r) \ n \ length vs" -using assms -apply(induct n arbitrary: v) -apply(erule Prf.cases) -apply(simp_all) -apply(erule Prf.cases) -apply(simp_all) -done - -lemma PLUS_Stars: - assumes "\ v : PLUS r" - shows "\vs. v = Stars vs \ (\v \ set vs. \ v : r) \ 1 \ length vs" -using assms -apply(erule_tac Prf.cases) -apply(simp_all) -done - -lemma NMTIMES_Stars: - assumes "\ v : NMTIMES r n m" - shows "\vs. v = Stars vs \ (\v \ set vs. \ v : r) \ n \ length vs \ length vs \ m" -using assms -apply(induct n arbitrary: m v) -apply(erule Prf.cases) -apply(simp_all) -apply(erule Prf.cases) -apply(simp_all) -done - +oops lemma Posix_injval: - assumes "s \ (der c r) \ v" + assumes "s \ (der c r) \ v" shows "(c # s) \ r \ (injval r c v)" using assms proof(induct r arbitrary: s v rule: rexp.induct) @@ -1352,7 +315,57 @@ then show "(c # s) \ SEQ r1 r2 \ injval (SEQ r1 r2) c v" using not_nullable by simp qed next - case (STAR r) +case (UPNTIMES r n s v) + have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact + have "s \ der c (UPNTIMES r n) \ v" by fact + then consider + (cons) v1 vs s1 s2 where + "v = Seq v1 (Stars vs)" "s = s1 @ s2" + "s1 \ der c r \ v1" "s2 \ (UPNTIMES r (n - 1)) \ (Stars vs)" "0 < n" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (UPNTIMES r (n - 1)))" + (* here *) + apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits) + apply(erule Posix_elims) + apply(simp) + apply(subgoal_tac "\vss. v2 = Stars vss") + apply(clarify) + apply(drule_tac x="v1" in meta_spec) + apply(drule_tac x="vss" in meta_spec) + apply(drule_tac x="s1" in meta_spec) + apply(drule_tac x="s2" in meta_spec) + apply(simp add: der_correctness Der_def) + apply(erule Posix_elims) + apply(auto) + done + then show "(c # s) \ (UPNTIMES r n) \ injval (UPNTIMES r n) c v" + proof (cases) + case cons + have "s1 \ der c r \ v1" by fact + then have "(c # s1) \ r \ injval r c v1" using IH by simp + moreover + have "s2 \ (UPNTIMES r (n - 1)) \ Stars vs" by fact + moreover + have "(c # s1) \ r \ injval r c v1" by fact + then have "flat (injval r c v1) = (c # s1)" by (rule Posix1) + then have "flat (injval r c v1) \ []" by simp + moreover + have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (UPNTIMES r (n - 1)))" by fact + then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (UPNTIMES r (n - 1)))" + by (simp add: der_correctness Der_def) + ultimately + have "((c # s1) @ s2) \ UPNTIMES r n \ Stars (injval r c v1 # vs)" + thm Posix.intros + apply (rule_tac Posix.intros) + apply(simp_all) + apply(case_tac n) + apply(simp) + using Posix_elims(1) UPNTIMES.prems apply auto[1] + apply(simp) + done + then show "(c # s) \ UPNTIMES r n \ injval (UPNTIMES r n) c v" using cons by(simp) + qed + next + case (STAR r) have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact have "s \ der c (STAR r) \ v" by fact then consider @@ -1384,242 +397,146 @@ have "((c # s1) @ s2) \ STAR r \ Stars (injval r c v1 # vs)" by (rule Posix.intros) then show "(c # s) \ STAR r \ injval (STAR r) c v" using cons by(simp) qed -next - case (UPNTIMES r n) - have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact - have "s \ der c (UPNTIMES r n) \ v" by fact + next + case (NTIMES r n s v) + have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact + have "s \ der c (NTIMES r n) \ v" by fact then consider - (cons) m v1 vs s1 s2 where - "n = Suc m" + (cons) v1 vs s1 s2 where "v = Seq v1 (Stars vs)" "s = s1 @ s2" - "s1 \ der c r \ v1" "s2 \ (UPNTIMES r m) \ (Stars vs)" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (UPNTIMES r m))" - apply(case_tac n) - apply(simp) - using Posix_elims(1) apply blast - apply(simp) - apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) - by (metis Posix1a UPNTIMES_Stars) - then show "(c # s) \ UPNTIMES r n \ injval (UPNTIMES r n) c v" + "s1 \ der c r \ v1" "s2 \ (NTIMES r (n - 1)) \ (Stars vs)" "0 < n" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (NTIMES r (n - 1)))" + apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits) + apply(erule Posix_elims) + apply(simp) + apply(subgoal_tac "\vss. v2 = Stars vss") + apply(clarify) + apply(drule_tac x="v1" in meta_spec) + apply(drule_tac x="vss" in meta_spec) + apply(drule_tac x="s1" in meta_spec) + apply(drule_tac x="s2" in meta_spec) + apply(simp add: der_correctness Der_def) + apply(erule Posix_elims) + apply(auto) + done + then show "(c # s) \ (NTIMES r n) \ injval (NTIMES r n) c v" proof (cases) case cons - have "n = Suc m" by fact - moreover have "s1 \ der c r \ v1" by fact then have "(c # s1) \ r \ injval r c v1" using IH by simp moreover - have "s2 \ UPNTIMES r m \ Stars vs" by fact + have "s2 \ (NTIMES r (n - 1)) \ Stars vs" by fact moreover have "(c # s1) \ r \ injval r c v1" by fact then have "flat (injval r c v1) = (c # s1)" by (rule Posix1) then have "flat (injval r c v1) \ []" by simp moreover - have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (UPNTIMES r m))" by fact - then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (UPNTIMES r m))" + have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (NTIMES r (n - 1)))" by fact + then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (NTIMES r (n - 1)))" by (simp add: der_correctness Der_def) ultimately - have "((c # s1) @ s2) \ UPNTIMES r (Suc m) \ Stars (injval r c v1 # vs)" - apply(rule_tac Posix.intros(8)) - apply(simp_all) - done - then show "(c # s) \ UPNTIMES r n \ injval (UPNTIMES r n) c v" using cons by(simp) - qed -next - case (NTIMES r n) - have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact - have "s \ der c (NTIMES r n) \ v" by fact - then consider - (cons) m v1 vs s1 s2 where - "n = Suc m" - "v = Seq v1 (Stars vs)" "s = s1 @ s2" "flat v = [] \ flat (Stars vs) = []" - "s1 \ der c r \ v1" "s2 \ (NTIMES r m) \ (Stars vs)" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (NTIMES r m))" - apply(case_tac n) - apply(simp) - using Posix_elims(1) apply blast - apply(simp) - apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) - by (metis NTIMES_Stars Posix1a) - then show "(c # s) \ NTIMES r n \ injval (NTIMES r n) c v" - proof (cases) - case cons - have "n = Suc m" by fact - moreover - have "s1 \ der c r \ v1" by fact - then have "(c # s1) \ r \ injval r c v1" using IH by simp - moreover - have "flat v = [] \ flat (Stars vs) = []" by fact - moreover - have "s2 \ NTIMES r m \ Stars vs" by fact - moreover - have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (NTIMES r m))" by fact - then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (NTIMES r m))" - by (simp add: der_correctness Der_def) - ultimately - have "((c # s1) @ s2) \ NTIMES r (Suc m) \ Stars (injval r c v1 # vs)" - apply(rule_tac Posix.intros(10)) - apply(simp_all) - by (simp add: Posix1(2)) + have "((c # s1) @ s2) \ NTIMES r n \ Stars (injval r c v1 # vs)" + apply (rule_tac Posix.intros) + apply(simp_all) + apply(case_tac n) + apply(simp) + using Posix_elims(1) NTIMES.prems apply auto[1] + apply(simp) + done then show "(c # s) \ NTIMES r n \ injval (NTIMES r n) c v" using cons by(simp) - qed -next - case (FROMNTIMES r n) + qed + next + case (FROMNTIMES r n s v) have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact have "s \ der c (FROMNTIMES r n) \ v" by fact then consider - (null) v1 vs s1 s2 where - "n = 0" + (cons) v1 vs s1 s2 where "v = Seq v1 (Stars vs)" "s = s1 @ s2" - "s1 \ der c r \ v1" "s2 \ (FROMNTIMES r 0) \ (Stars vs)" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (FROMNTIMES r 0))" - | (cons) m v1 vs s1 s2 where - "n = Suc m" - "v = Seq v1 (Stars vs)" "s = s1 @ s2" "flat v = [] \ flat (Stars vs) = []" - "s1 \ der c r \ v1" "s2 \ (FROMNTIMES r m) \ (Stars vs)" - "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (FROMNTIMES r m))" - apply(case_tac n) - apply(simp) - apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) - defer - apply (metis FROMNTIMES_Stars Posix1a) - apply(rotate_tac 5) - apply(erule Posix.cases) - apply(simp_all) - apply(clarify) - by (simp add: Posix_FROMNTIMES2) - then show "(c # s) \ FROMNTIMES r n \ injval (FROMNTIMES r n) c v" + "s1 \ der c r \ v1" "s2 \ (FROMNTIMES r (n - 1)) \ (Stars vs)" "0 < n" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (FROMNTIMES r (n - 1)))" + | (null) v1 where + "v = Seq v1 (Stars [])" + "s \ der c r \ v1" "[] \ (FROMNTIMES r 0) \ (Stars [])" "n = 0" + (* here *) + apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits) + apply(erule Posix_elims) + apply(simp) + apply(case_tac "n = 0") + apply(simp) + apply(drule FROMNTIMES_0) + apply(simp) + using FROMNTIMES_0 Posix_mkeps apply force + apply(subgoal_tac "\vss. v2 = Stars vss") + apply(clarify) + apply(drule_tac x="v1" in meta_spec) + apply(drule_tac x="vss" in meta_spec) + apply(drule_tac x="s1" in meta_spec) + apply(drule_tac x="s2" in meta_spec) + apply(simp add: der_correctness Der_def) + apply(case_tac "n = 0") + apply(simp) + apply(simp) + apply(rotate_tac 4) + apply(erule Posix_elims) + apply(auto)[2] + done + then show "(c # s) \ (FROMNTIMES r n) \ injval (FROMNTIMES r n) c v" proof (cases) case cons - have "n = Suc m" by fact - moreover have "s1 \ der c r \ v1" by fact then have "(c # s1) \ r \ injval r c v1" using IH by simp moreover - have "s2 \ FROMNTIMES r m \ Stars vs" by fact - moreover - have "flat v = [] \ flat (Stars vs) = []" by fact + have "s2 \ (FROMNTIMES r (n - 1)) \ Stars vs" by fact moreover - have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (FROMNTIMES r m))" by fact - then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (FROMNTIMES r m))" + have "(c # s1) \ r \ injval r c v1" by fact + then have "flat (injval r c v1) = (c # s1)" by (rule Posix1) + then have "flat (injval r c v1) \ []" by simp + moreover + have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (FROMNTIMES r (n - 1)))" by fact + then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (FROMNTIMES r (n - 1)))" by (simp add: der_correctness Der_def) ultimately - have "((c # s1) @ s2) \ FROMNTIMES r (Suc m) \ Stars (injval r c v1 # vs)" - apply(rule_tac Posix.intros(12)) - apply(simp_all) - by (simp add: Posix1(2)) + have "((c # s1) @ s2) \ FROMNTIMES r n \ Stars (injval r c v1 # vs)" + apply (rule_tac Posix.intros) + apply(simp_all) + apply(case_tac n) + apply(simp) + using Posix_elims(1) FROMNTIMES.prems apply auto[1] + using cons(5) apply blast + apply(simp) + done then show "(c # s) \ FROMNTIMES r n \ injval (FROMNTIMES r n) c v" using cons by(simp) - next - case null - then have "((c # s1) @ s2) \ FROMNTIMES r 0 \ Stars (injval r c v1 # vs)" - apply(rule_tac Posix.intros) - apply(rule_tac Posix.intros) - apply(rule IH) - apply(simp) - apply(rotate_tac 4) - apply(erule Posix.cases) - apply(simp_all) - apply (simp add: Posix1a Prf_injval_flat) - apply(simp only: Star_def) - apply(simp only: der_correctness Der_def) - apply(simp) - done - then show "(c # s) \ FROMNTIMES r n \ injval (FROMNTIMES r n) c v" using null by simp - qed -next - case (NMTIMES r n m) - have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact - have "s \ der c (NMTIMES r n m) \ v" by fact - then show "(c # s) \ (NMTIMES r n m) \ injval (NMTIMES r n m) c v" - apply(case_tac "m < n") - apply(simp) - using Posix_elims(1) apply blast - apply(case_tac n) - apply(simp_all) - apply(case_tac m) - apply(simp) - apply(erule Posix_elims(1)) - apply(simp) - apply(erule Posix.cases) - apply(simp_all) - apply(clarify) - apply(rotate_tac 5) - apply(frule Posix1a) - apply(drule UPNTIMES_Stars) - apply(clarify) - apply(simp) - apply(subst append_Cons[symmetric]) - apply(rule Posix.intros) - apply(rule Posix.intros) - apply(auto) - apply(rule IH) - apply blast - using Posix1a Prf_injval_flat apply auto[1] - apply(simp add: der_correctness Der_def) - apply blast - apply(case_tac m) - apply(simp) - apply(simp) - apply(erule Posix.cases) - apply(simp_all) - apply(clarify) - apply(rotate_tac 6) - apply(frule Posix1a) - apply(drule NMTIMES_Stars) - apply(clarify) - apply(simp) - apply(subst append_Cons[symmetric]) - apply(rule Posix.intros) - apply(rule IH) - apply(blast) - apply(simp) - apply(simp add: der_correctness Der_def) - using Posix1a Prf_injval_flat list.distinct(1) apply auto[1] - apply(simp add: der_correctness Der_def) - done -next - case (PLUS r) - have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact - have "s \ der c (PLUS r) \ v" by fact - then show "(c # s) \ PLUS r \ injval (PLUS r) c v" - apply - - apply(erule Posix.cases) - apply(simp_all) - apply(auto) - apply(rotate_tac 3) - apply(erule Posix.cases) - apply(simp_all) - apply(subst append_Cons[symmetric]) - apply(rule Posix.intros) - using PLUS.hyps apply auto[1] - apply(rule Posix.intros) - apply(simp) - apply(simp) - apply(simp) - apply(simp) - apply(simp) - using Posix1a Prf_injval_flat apply auto[1] - apply(simp add: der_correctness Der_def) - apply(subst append_Nil2[symmetric]) - apply(rule Posix.intros) - using PLUS.hyps apply auto[1] - apply(rule Posix.intros) - apply(simp) - apply(simp) - done + next + case null + have "s \ der c r \ v1" by fact + then have "(c # s) \ r \ injval r c v1" using IH by simp + moreover + have "[] \ (FROMNTIMES r 0) \ Stars []" by fact + moreover + have "(c # s) \ r \ injval r c v1" by fact + then have "flat (injval r c v1) = (c # s)" by (rule Posix1) + then have "flat (injval r c v1) \ []" by simp + ultimately + have "((c # s) @ []) \ FROMNTIMES r 1 \ Stars [injval r c v1]" + apply (rule_tac Posix.intros) + apply(simp_all) + done + then show "(c # s) \ FROMNTIMES r n \ injval (FROMNTIMES r n) c v" using null + apply(simp) + apply(erule Posix_elims) + apply(auto) + apply(rotate_tac 6) + apply(drule FROMNTIMES_0) + apply(simp) + sorry + qed + next + case (NMTIMES x1 x2 m s v) + then show ?case sorry qed - -section {* The Lexer by Sulzmann and Lu *} - -fun - lexer :: "rexp \ string \ val option" -where - "lexer r [] = (if nullable r then Some(mkeps r) else None)" -| "lexer r (c#s) = (case (lexer (der c r) s) of - None \ None - | Some(v) \ Some(injval r c v))" - +section {* Lexer Correctness *} lemma lexer_correct_None: shows "s \ L r \ lexer r s = None" @@ -1645,8 +562,4 @@ using Posix1(1) Posix_determ lexer_correct_None lexer_correct_Some apply fastforce using Posix1(1) lexer_correct_None lexer_correct_Some by blast -value "lexer (PLUS (ALT ONE (SEQ (CHAR (CHR ''a'')) (CHAR (CHR ''b''))))) [CHR ''a'', CHR ''b'']" - - -unused_thms end \ No newline at end of file diff -r deea42c83c9e -r a3134f7de065 thys/PositionsExt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/PositionsExt.thy Sat Oct 07 22:16:16 2017 +0100 @@ -0,0 +1,990 @@ + +theory PositionsExt + imports "SpecExt" "LexerExt" +begin + +section {* Positions in Values *} + +fun + at :: "val \ nat list \ val" +where + "at v [] = v" +| "at (Left v) (0#ps)= at v ps" +| "at (Right v) (Suc 0#ps)= at v ps" +| "at (Seq v1 v2) (0#ps)= at v1 ps" +| "at (Seq v1 v2) (Suc 0#ps)= at v2 ps" +| "at (Stars vs) (n#ps)= at (nth vs n) ps" + + + +fun Pos :: "val \ (nat list) set" +where + "Pos (Void) = {[]}" +| "Pos (Char c) = {[]}" +| "Pos (Left v) = {[]} \ {0#ps | ps. ps \ Pos v}" +| "Pos (Right v) = {[]} \ {1#ps | ps. ps \ Pos v}" +| "Pos (Seq v1 v2) = {[]} \ {0#ps | ps. ps \ Pos v1} \ {1#ps | ps. ps \ Pos v2}" +| "Pos (Stars []) = {[]}" +| "Pos (Stars (v#vs)) = {[]} \ {0#ps | ps. ps \ Pos v} \ {Suc n#ps | n ps. n#ps \ Pos (Stars vs)}" + + +lemma Pos_stars: + "Pos (Stars vs) = {[]} \ (\n < length vs. {n#ps | ps. ps \ Pos (vs ! n)})" +apply(induct vs) +apply(auto simp add: insert_ident less_Suc_eq_0_disj) +done + +lemma Pos_empty: + shows "[] \ Pos v" +by (induct v rule: Pos.induct)(auto) + + +abbreviation + "intlen vs \ int (length vs)" + + +definition pflat_len :: "val \ nat list => int" +where + "pflat_len v p \ (if p \ Pos v then intlen (flat (at v p)) else -1)" + +lemma pflat_len_simps: + shows "pflat_len (Seq v1 v2) (0#p) = pflat_len v1 p" + and "pflat_len (Seq v1 v2) (Suc 0#p) = pflat_len v2 p" + and "pflat_len (Left v) (0#p) = pflat_len v p" + and "pflat_len (Left v) (Suc 0#p) = -1" + and "pflat_len (Right v) (Suc 0#p) = pflat_len v p" + and "pflat_len (Right v) (0#p) = -1" + and "pflat_len (Stars (v#vs)) (Suc n#p) = pflat_len (Stars vs) (n#p)" + and "pflat_len (Stars (v#vs)) (0#p) = pflat_len v p" + and "pflat_len v [] = intlen (flat v)" +by (auto simp add: pflat_len_def Pos_empty) + +lemma pflat_len_Stars_simps: + assumes "n < length vs" + shows "pflat_len (Stars vs) (n#p) = pflat_len (vs!n) p" +using assms +apply(induct vs arbitrary: n p) +apply(auto simp add: less_Suc_eq_0_disj pflat_len_simps) +done + +lemma pflat_len_outside: + assumes "p \ Pos v1" + shows "pflat_len v1 p = -1 " +using assms by (simp add: pflat_len_def) + + + +section {* Orderings *} + + +definition prefix_list:: "'a list \ 'a list \ bool" ("_ \pre _" [60,59] 60) +where + "ps1 \pre ps2 \ \ps'. ps1 @ps' = ps2" + +definition sprefix_list:: "'a list \ 'a list \ bool" ("_ \spre _" [60,59] 60) +where + "ps1 \spre ps2 \ ps1 \pre ps2 \ ps1 \ ps2" + +inductive lex_list :: "nat list \ nat list \ bool" ("_ \lex _" [60,59] 60) +where + "[] \lex (p#ps)" +| "ps1 \lex ps2 \ (p#ps1) \lex (p#ps2)" +| "p1 < p2 \ (p1#ps1) \lex (p2#ps2)" + +lemma lex_irrfl: + fixes ps1 ps2 :: "nat list" + assumes "ps1 \lex ps2" + shows "ps1 \ ps2" +using assms +by(induct rule: lex_list.induct)(auto) + +lemma lex_simps [simp]: + fixes xs ys :: "nat list" + shows "[] \lex ys \ ys \ []" + and "xs \lex [] \ False" + and "(x # xs) \lex (y # ys) \ (x < y \ (x = y \ xs \lex ys))" +by (auto simp add: neq_Nil_conv elim: lex_list.cases intro: lex_list.intros) + +lemma lex_trans: + fixes ps1 ps2 ps3 :: "nat list" + assumes "ps1 \lex ps2" "ps2 \lex ps3" + shows "ps1 \lex ps3" +using assms +by (induct arbitrary: ps3 rule: lex_list.induct) + (auto elim: lex_list.cases) + + +lemma lex_trichotomous: + fixes p q :: "nat list" + shows "p = q \ p \lex q \ q \lex p" +apply(induct p arbitrary: q) +apply(auto elim: lex_list.cases) +apply(case_tac q) +apply(auto) +done + + + + +section {* POSIX Ordering of Values According to Okui & Suzuki *} + + +definition PosOrd:: "val \ nat list \ val \ bool" ("_ \val _ _" [60, 60, 59] 60) +where + "v1 \val p v2 \ pflat_len v1 p > pflat_len v2 p \ + (\q \ Pos v1 \ Pos v2. q \lex p \ pflat_len v1 q = pflat_len v2 q)" + +lemma PosOrd_def2: + shows "v1 \val p v2 \ + pflat_len v1 p > pflat_len v2 p \ + (\q \ Pos v1. q \lex p \ pflat_len v1 q = pflat_len v2 q) \ + (\q \ Pos v2. q \lex p \ pflat_len v1 q = pflat_len v2 q)" +unfolding PosOrd_def +apply(auto) +done + + +definition PosOrd_ex:: "val \ val \ bool" ("_ :\val _" [60, 59] 60) +where + "v1 :\val v2 \ \p. v1 \val p v2" + +definition PosOrd_ex_eq:: "val \ val \ bool" ("_ :\val _" [60, 59] 60) +where + "v1 :\val v2 \ v1 :\val v2 \ v1 = v2" + + +lemma PosOrd_trans: + assumes "v1 :\val v2" "v2 :\val v3" + shows "v1 :\val v3" +proof - + from assms obtain p p' + where as: "v1 \val p v2" "v2 \val p' v3" unfolding PosOrd_ex_def by blast + then have pos: "p \ Pos v1" "p' \ Pos v2" unfolding PosOrd_def pflat_len_def + by (smt not_int_zless_negative)+ + have "p = p' \ p \lex p' \ p' \lex p" + by (rule lex_trichotomous) + moreover + { assume "p = p'" + with as have "v1 \val p v3" unfolding PosOrd_def pflat_len_def + by (smt Un_iff) + then have " v1 :\val v3" unfolding PosOrd_ex_def by blast + } + moreover + { assume "p \lex p'" + with as have "v1 \val p v3" unfolding PosOrd_def pflat_len_def + by (smt Un_iff lex_trans) + then have " v1 :\val v3" unfolding PosOrd_ex_def by blast + } + moreover + { assume "p' \lex p" + with as have "v1 \val p' v3" unfolding PosOrd_def + by (smt Un_iff lex_trans pflat_len_def) + then have "v1 :\val v3" unfolding PosOrd_ex_def by blast + } + ultimately show "v1 :\val v3" by blast +qed + +lemma PosOrd_irrefl: + assumes "v :\val v" + shows "False" +using assms unfolding PosOrd_ex_def PosOrd_def +by auto + +lemma PosOrd_assym: + assumes "v1 :\val v2" + shows "\(v2 :\val v1)" +using assms +using PosOrd_irrefl PosOrd_trans by blast + +text {* + :\val and :\val are partial orders. +*} + +lemma PosOrd_ordering: + shows "ordering (\v1 v2. v1 :\val v2) (\ v1 v2. v1 :\val v2)" +unfolding ordering_def PosOrd_ex_eq_def +apply(auto) +using PosOrd_irrefl apply blast +using PosOrd_assym apply blast +using PosOrd_trans by blast + +lemma PosOrd_order: + shows "class.order (\v1 v2. v1 :\val v2) (\ v1 v2. v1 :\val v2)" +using PosOrd_ordering +apply(simp add: class.order_def class.preorder_def class.order_axioms_def) +unfolding ordering_def +by blast + + +lemma PosOrd_ex_eq2: + shows "v1 :\val v2 \ (v1 :\val v2 \ v1 \ v2)" +using PosOrd_ordering +unfolding ordering_def +by auto + +lemma PosOrdeq_trans: + assumes "v1 :\val v2" "v2 :\val v3" + shows "v1 :\val v3" +using assms PosOrd_ordering +unfolding ordering_def +by blast + +lemma PosOrdeq_antisym: + assumes "v1 :\val v2" "v2 :\val v1" + shows "v1 = v2" +using assms PosOrd_ordering +unfolding ordering_def +by blast + +lemma PosOrdeq_refl: + shows "v :\val v" +unfolding PosOrd_ex_eq_def +by auto + + +lemma PosOrd_shorterE: + assumes "v1 :\val v2" + shows "length (flat v2) \ length (flat v1)" +using assms unfolding PosOrd_ex_def PosOrd_def +apply(auto) +apply(case_tac p) +apply(simp add: pflat_len_simps) +apply(drule_tac x="[]" in bspec) +apply(simp add: Pos_empty) +apply(simp add: pflat_len_simps) +done + +lemma PosOrd_shorterI: + assumes "length (flat v2) < length (flat v1)" + shows "v1 :\val v2" +unfolding PosOrd_ex_def PosOrd_def pflat_len_def +using assms Pos_empty by force + +lemma PosOrd_spreI: + assumes "flat v' \spre flat v" + shows "v :\val v'" +using assms +apply(rule_tac PosOrd_shorterI) +unfolding prefix_list_def sprefix_list_def +by (metis append_Nil2 append_eq_conv_conj drop_all le_less_linear) + +lemma pflat_len_inside: + assumes "pflat_len v2 p < pflat_len v1 p" + shows "p \ Pos v1" +using assms +unfolding pflat_len_def +by (auto split: if_splits) + + +lemma PosOrd_Left_Right: + assumes "flat v1 = flat v2" + shows "Left v1 :\val Right v2" +unfolding PosOrd_ex_def +apply(rule_tac x="[0]" in exI) +apply(auto simp add: PosOrd_def pflat_len_simps assms) +done + +lemma PosOrd_LeftE: + assumes "Left v1 :\val Left v2" "flat v1 = flat v2" + shows "v1 :\val v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +apply(frule pflat_len_inside) +apply(auto simp add: pflat_len_simps) +by (metis lex_simps(3) pflat_len_simps(3)) + +lemma PosOrd_LeftI: + assumes "v1 :\val v2" "flat v1 = flat v2" + shows "Left v1 :\val Left v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +by (metis less_numeral_extra(3) lex_simps(3) pflat_len_simps(3)) + +lemma PosOrd_Left_eq: + assumes "flat v1 = flat v2" + shows "Left v1 :\val Left v2 \ v1 :\val v2" +using assms PosOrd_LeftE PosOrd_LeftI +by blast + + +lemma PosOrd_RightE: + assumes "Right v1 :\val Right v2" "flat v1 = flat v2" + shows "v1 :\val v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +apply(frule pflat_len_inside) +apply(auto simp add: pflat_len_simps) +by (metis lex_simps(3) pflat_len_simps(5)) + +lemma PosOrd_RightI: + assumes "v1 :\val v2" "flat v1 = flat v2" + shows "Right v1 :\val Right v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +by (metis lex_simps(3) nat_neq_iff pflat_len_simps(5)) + + +lemma PosOrd_Right_eq: + assumes "flat v1 = flat v2" + shows "Right v1 :\val Right v2 \ v1 :\val v2" +using assms PosOrd_RightE PosOrd_RightI +by blast + + +lemma PosOrd_SeqI1: + assumes "v1 :\val w1" "flat (Seq v1 v2) = flat (Seq w1 w2)" + shows "Seq v1 v2 :\val Seq w1 w2" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(rule_tac x="0#p" in exI) +apply(subst PosOrd_def) +apply(rule conjI) +apply(simp add: pflat_len_simps) +apply(rule ballI) +apply(rule impI) +apply(simp only: Pos.simps) +apply(auto)[1] +apply(simp add: pflat_len_simps) +apply(auto simp add: pflat_len_simps) +using assms(2) +apply(simp) +apply(metis length_append of_nat_add) +done + +lemma PosOrd_SeqI2: + assumes "v2 :\val w2" "flat v2 = flat w2" + shows "Seq v v2 :\val Seq v w2" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(rule_tac x="Suc 0#p" in exI) +apply(subst PosOrd_def) +apply(rule conjI) +apply(simp add: pflat_len_simps) +apply(rule ballI) +apply(rule impI) +apply(simp only: Pos.simps) +apply(auto)[1] +apply(simp add: pflat_len_simps) +using assms(2) +apply(simp) +apply(auto simp add: pflat_len_simps) +done + +lemma PosOrd_Seq_eq: + assumes "flat v2 = flat w2" + shows "(Seq v v2) :\val (Seq v w2) \ v2 :\val w2" +using assms +apply(auto) +prefer 2 +apply(simp add: PosOrd_SeqI2) +apply(simp add: PosOrd_ex_def) +apply(auto) +apply(case_tac p) +apply(simp add: PosOrd_def pflat_len_simps) +apply(case_tac a) +apply(simp add: PosOrd_def pflat_len_simps) +apply(clarify) +apply(case_tac nat) +prefer 2 +apply(simp add: PosOrd_def pflat_len_simps pflat_len_outside) +apply(rule_tac x="list" in exI) +apply(auto simp add: PosOrd_def2 pflat_len_simps) +apply(smt Collect_disj_eq lex_list.intros(2) mem_Collect_eq pflat_len_simps(2)) +apply(smt Collect_disj_eq lex_list.intros(2) mem_Collect_eq pflat_len_simps(2)) +done + + + +lemma PosOrd_StarsI: + assumes "v1 :\val v2" "flats (v1#vs1) = flats (v2#vs2)" + shows "Stars (v1#vs1) :\val Stars (v2#vs2)" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(subst PosOrd_def) +apply(rule_tac x="0#p" in exI) +apply(simp add: pflat_len_Stars_simps pflat_len_simps) +using assms(2) +apply(simp add: pflat_len_simps) +apply(auto simp add: pflat_len_Stars_simps pflat_len_simps) +by (metis length_append of_nat_add) + +lemma PosOrd_StarsI2: + assumes "Stars vs1 :\val Stars vs2" "flats vs1 = flats vs2" + shows "Stars (v#vs1) :\val Stars (v#vs2)" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(subst PosOrd_def) +apply(case_tac p) +apply(simp add: pflat_len_simps) +apply(rule_tac x="Suc a#list" in exI) +apply(auto simp add: pflat_len_Stars_simps pflat_len_simps assms(2)) +done + +lemma PosOrd_Stars_appendI: + assumes "Stars vs1 :\val Stars vs2" "flat (Stars vs1) = flat (Stars vs2)" + shows "Stars (vs @ vs1) :\val Stars (vs @ vs2)" +using assms +apply(induct vs) +apply(simp) +apply(simp add: PosOrd_StarsI2) +done + +lemma PosOrd_eq_Stars_zipI: + assumes "\(v1, v2) \ set (zip vs1 vs2). v1 :\val v2" + "length vs1 = length vs2" "flats vs1 = flats vs2" + shows "Stars vs1 :\val Stars vs2" + using assms + apply(induct vs1 arbitrary: vs2) + apply(case_tac vs2) +apply(simp add: PosOrd_ex_eq_def) + apply(simp) + apply(case_tac vs2) + apply(simp) + apply(simp) + apply(auto) +apply(subst (asm) (2)PosOrd_ex_eq_def) + apply(auto) + apply(subst PosOrd_ex_eq_def) + apply(rule disjI1) + apply(rule PosOrd_StarsI) + apply(simp) + apply(simp) + using PosOrd_StarsI2 PosOrd_ex_eq_def by fastforce + +lemma PosOrd_StarsE2: + assumes "Stars (v # vs1) :\val Stars (v # vs2)" + shows "Stars vs1 :\val Stars vs2" +using assms +apply(subst (asm) PosOrd_ex_def) +apply(erule exE) +apply(case_tac p) +apply(simp) +apply(simp add: PosOrd_def pflat_len_simps) +apply(subst PosOrd_ex_def) +apply(rule_tac x="[]" in exI) +apply(simp add: PosOrd_def pflat_len_simps Pos_empty) +apply(simp) +apply(case_tac a) +apply(clarify) +apply(auto simp add: pflat_len_simps PosOrd_def pflat_len_def split: if_splits)[1] +apply(clarify) +apply(simp add: PosOrd_ex_def) +apply(rule_tac x="nat#list" in exI) +apply(auto simp add: PosOrd_def pflat_len_simps)[1] +apply(case_tac q) +apply(simp add: PosOrd_def pflat_len_simps) +apply(clarify) +apply(drule_tac x="Suc a # lista" in bspec) +apply(simp) +apply(auto simp add: PosOrd_def pflat_len_simps)[1] +apply(case_tac q) +apply(simp add: PosOrd_def pflat_len_simps) +apply(clarify) +apply(drule_tac x="Suc a # lista" in bspec) +apply(simp) +apply(auto simp add: PosOrd_def pflat_len_simps)[1] +done + +lemma PosOrd_Stars_appendE: + assumes "Stars (vs @ vs1) :\val Stars (vs @ vs2)" + shows "Stars vs1 :\val Stars vs2" +using assms +apply(induct vs) +apply(simp) +apply(simp add: PosOrd_StarsE2) +done + +lemma PosOrd_Stars_append_eq: + assumes "flats vs1 = flats vs2" + shows "Stars (vs @ vs1) :\val Stars (vs @ vs2) \ Stars vs1 :\val Stars vs2" +using assms +apply(rule_tac iffI) +apply(erule PosOrd_Stars_appendE) +apply(rule PosOrd_Stars_appendI) +apply(auto) +done + +lemma PosOrd_almost_trichotomous: + shows "v1 :\val v2 \ v2 :\val v1 \ (length (flat v1) = length (flat v2))" +apply(auto simp add: PosOrd_ex_def) +apply(auto simp add: PosOrd_def) +apply(rule_tac x="[]" in exI) +apply(auto simp add: Pos_empty pflat_len_simps) +apply(drule_tac x="[]" in spec) +apply(auto simp add: Pos_empty pflat_len_simps) +done + + + +section {* The Posix Value is smaller than any other Value *} + + +lemma Posix_PosOrd: + assumes "s \ r \ v1" "v2 \ LV r s" + shows "v1 :\val v2" +using assms +proof (induct arbitrary: v2 rule: Posix.induct) + case (Posix_ONE v) + have "v \ LV ONE []" by fact + then have "v = Void" + by (simp add: LV_simps) + then show "Void :\val v" + by (simp add: PosOrd_ex_eq_def) +next + case (Posix_CHAR c v) + have "v \ LV (CHAR c) [c]" by fact + then have "v = Char c" + by (simp add: LV_simps) + then show "Char c :\val v" + by (simp add: PosOrd_ex_eq_def) +next + case (Posix_ALT1 s r1 v r2 v2) + have as1: "s \ r1 \ v" by fact + have IH: "\v2. v2 \ LV r1 s \ v :\val v2" by fact + have "v2 \ LV (ALT r1 r2) s" by fact + then have "\ v2 : ALT r1 r2" "flat v2 = s" + by(auto simp add: LV_def prefix_list_def) + then consider + (Left) v3 where "v2 = Left v3" "\ v3 : r1" "flat v3 = s" + | (Right) v3 where "v2 = Right v3" "\ v3 : r2" "flat v3 = s" + by (auto elim: Prf.cases) + then show "Left v :\val v2" + proof(cases) + case (Left v3) + have "v3 \ LV r1 s" using Left(2,3) + by (auto simp add: LV_def prefix_list_def) + with IH have "v :\val v3" by simp + moreover + have "flat v3 = flat v" using as1 Left(3) + by (simp add: Posix1(2)) + ultimately have "Left v :\val Left v3" + by (simp add: PosOrd_ex_eq_def PosOrd_Left_eq) + then show "Left v :\val v2" unfolding Left . + next + case (Right v3) + have "flat v3 = flat v" using as1 Right(3) + by (simp add: Posix1(2)) + then have "Left v :\val Right v3" + unfolding PosOrd_ex_eq_def + by (simp add: PosOrd_Left_Right) + then show "Left v :\val v2" unfolding Right . + qed +next + case (Posix_ALT2 s r2 v r1 v2) + have as1: "s \ r2 \ v" by fact + have as2: "s \ L r1" by fact + have IH: "\v2. v2 \ LV r2 s \ v :\val v2" by fact + have "v2 \ LV (ALT r1 r2) s" by fact + then have "\ v2 : ALT r1 r2" "flat v2 = s" + by(auto simp add: LV_def prefix_list_def) + then consider + (Left) v3 where "v2 = Left v3" "\ v3 : r1" "flat v3 = s" + | (Right) v3 where "v2 = Right v3" "\ v3 : r2" "flat v3 = s" + by (auto elim: Prf.cases) + then show "Right v :\val v2" + proof (cases) + case (Right v3) + have "v3 \ LV r2 s" using Right(2,3) + by (auto simp add: LV_def prefix_list_def) + with IH have "v :\val v3" by simp + moreover + have "flat v3 = flat v" using as1 Right(3) + by (simp add: Posix1(2)) + ultimately have "Right v :\val Right v3" + by (auto simp add: PosOrd_ex_eq_def PosOrd_RightI) + then show "Right v :\val v2" unfolding Right . + next + case (Left v3) + have "v3 \ LV r1 s" using Left(2,3) as2 + by (auto simp add: LV_def prefix_list_def) + then have "flat v3 = flat v \ \ v3 : r1" using as1 Left(3) + by (simp add: Posix1(2) LV_def) + then have "False" using as1 as2 Left + by (auto simp add: Posix1(2) L_flat_Prf1) + then show "Right v :\val v2" by simp + qed +next + case (Posix_SEQ s1 r1 v1 s2 r2 v2 v3) + have "s1 \ r1 \ v1" "s2 \ r2 \ v2" by fact+ + then have as1: "s1 = flat v1" "s2 = flat v2" by (simp_all add: Posix1(2)) + have IH1: "\v3. v3 \ LV r1 s1 \ v1 :\val v3" by fact + have IH2: "\v3. v3 \ LV r2 s2 \ v2 :\val v3" by fact + have cond: "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" by fact + have "v3 \ LV (SEQ r1 r2) (s1 @ s2)" by fact + then obtain v3a v3b where eqs: + "v3 = Seq v3a v3b" "\ v3a : r1" "\ v3b : r2" + "flat v3a @ flat v3b = s1 @ s2" + by (force simp add: prefix_list_def LV_def elim: Prf.cases) + with cond have "flat v3a \pre s1" unfolding prefix_list_def + by (smt L_flat_Prf1 append_eq_append_conv2 append_self_conv) + then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat v3b = s2)" using eqs + by (simp add: sprefix_list_def append_eq_conv_conj) + then have q2: "v1 :\val v3a \ (flat v3a = s1 \ flat v3b = s2)" + using PosOrd_spreI as1(1) eqs by blast + then have "v1 :\val v3a \ (v3a \ LV r1 s1 \ v3b \ LV r2 s2)" using eqs(2,3) + by (auto simp add: LV_def) + then have "v1 :\val v3a \ (v1 :\val v3a \ v2 :\val v3b)" using IH1 IH2 by blast + then have "Seq v1 v2 :\val Seq v3a v3b" using eqs q2 as1 + unfolding PosOrd_ex_eq_def by (auto simp add: PosOrd_SeqI1 PosOrd_Seq_eq) + then show "Seq v1 v2 :\val v3" unfolding eqs by blast +next + case (Posix_STAR1 s1 r v s2 vs v3) + have "s1 \ r \ v" "s2 \ STAR r \ Stars vs" by fact+ + then have as1: "s1 = flat v" "s2 = flat (Stars vs)" by (auto dest: Posix1(2)) + have IH1: "\v3. v3 \ LV r s1 \ v :\val v3" by fact + have IH2: "\v3. v3 \ LV (STAR r) s2 \ Stars vs :\val v3" by fact + have cond: "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (STAR r))" by fact + have cond2: "flat v \ []" by fact + have "v3 \ LV (STAR r) (s1 @ s2)" by fact + then consider + (NonEmpty) v3a vs3 where "v3 = Stars (v3a # vs3)" + "\ v3a : r" "\ Stars vs3 : STAR r" + "flat (Stars (v3a # vs3)) = s1 @ s2" + | (Empty) "v3 = Stars []" + unfolding LV_def + apply(auto) + apply(erule Prf.cases) + apply(auto) + apply(case_tac vs) + apply(auto intro: Prf.intros) + done + then show "Stars (v # vs) :\val v3" + proof (cases) + case (NonEmpty v3a vs3) + have "flat (Stars (v3a # vs3)) = s1 @ s2" using NonEmpty(4) . + with cond have "flat v3a \pre s1" using NonEmpty(2,3) + unfolding prefix_list_def + by (smt L_flat_Prf1 append_Nil2 append_eq_append_conv2 flat.simps(7)) + then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat (Stars vs3) = s2)" using NonEmpty(4) + by (simp add: sprefix_list_def append_eq_conv_conj) + then have q2: "v :\val v3a \ (flat v3a = s1 \ flat (Stars vs3) = s2)" + using PosOrd_spreI as1(1) NonEmpty(4) by blast + then have "v :\val v3a \ (v3a \ LV r s1 \ Stars vs3 \ LV (STAR r) s2)" + using NonEmpty(2,3) by (auto simp add: LV_def) + then have "v :\val v3a \ (v :\val v3a \ Stars vs :\val Stars vs3)" using IH1 IH2 by blast + then have "v :\val v3a \ (v = v3a \ Stars vs :\val Stars vs3)" + unfolding PosOrd_ex_eq_def by auto + then have "Stars (v # vs) :\val Stars (v3a # vs3)" using NonEmpty(4) q2 as1 + unfolding PosOrd_ex_eq_def + using PosOrd_StarsI PosOrd_StarsI2 by auto + then show "Stars (v # vs) :\val v3" unfolding NonEmpty by blast + next + case Empty + have "v3 = Stars []" by fact + then show "Stars (v # vs) :\val v3" + unfolding PosOrd_ex_eq_def using cond2 + by (simp add: PosOrd_shorterI) + qed +next + case (Posix_STAR2 r v2) + have "v2 \ LV (STAR r) []" by fact + then have "v2 = Stars []" + unfolding LV_def by (auto elim: Prf.cases) + then show "Stars [] :\val v2" + by (simp add: PosOrd_ex_eq_def) +next + case (Posix_NTIMES1 s1 r v s2 n vs v3) + have "s1 \ r \ v" "s2 \ NTIMES r (n - 1) \ Stars vs" by fact+ + then have as1: "s1 = flat v" "s2 = flats vs" by (auto dest: Posix1(2)) + have IH1: "\v3. v3 \ LV r s1 \ v :\val v3" by fact + have IH2: "\v3. v3 \ LV (NTIMES r (n - 1)) s2 \ Stars vs :\val v3" by fact + have cond: "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (NTIMES r (n - 1)))" by fact + have cond2: "flat v \ []" by fact + have "v3 \ LV (NTIMES r n) (s1 @ s2)" by fact + then consider + (NonEmpty) v3a vs3 where "v3 = Stars (v3a # vs3)" + "\ v3a : r" "\ Stars vs3 : NTIMES r (n - 1)" + "flats (v3a # vs3) = s1 @ s2" + | (Empty) "v3 = Stars []" + unfolding LV_def + apply(auto) + apply(erule Prf.cases) + apply(auto) + apply(case_tac vs1) + apply(auto intro: Prf.intros) + apply(case_tac vs2) + apply(auto intro: Prf.intros) + apply (simp add: as1(1) cond2 flats_empty) + by (simp add: Prf.intros(8)) + then show "Stars (v # vs) :\val v3" + proof (cases) + case (NonEmpty v3a vs3) + have "flats (v3a # vs3) = s1 @ s2" using NonEmpty(4) . + with cond have "flat v3a \pre s1" using NonEmpty(2,3) + unfolding prefix_list_def + by (smt L_flat_Prf1 append_Nil2 append_eq_append_conv2 flat.simps(7) flat_Stars) + then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat (Stars vs3) = s2)" using NonEmpty(4) + by (simp add: sprefix_list_def append_eq_conv_conj) + then have q2: "v :\val v3a \ (flat v3a = s1 \ flat (Stars vs3) = s2)" + using PosOrd_spreI as1(1) NonEmpty(4) by blast + then have "v :\val v3a \ (v3a \ LV r s1 \ Stars vs3 \ LV (NTIMES r (n - 1)) s2)" + using NonEmpty(2,3) by (auto simp add: LV_def) + then have "v :\val v3a \ (v :\val v3a \ Stars vs :\val Stars vs3)" using IH1 IH2 by blast + then have "v :\val v3a \ (v = v3a \ Stars vs :\val Stars vs3)" + unfolding PosOrd_ex_eq_def by auto + then have "Stars (v # vs) :\val Stars (v3a # vs3)" using NonEmpty(4) q2 as1 + unfolding PosOrd_ex_eq_def + using PosOrd_StarsI PosOrd_StarsI2 by auto + then show "Stars (v # vs) :\val v3" unfolding NonEmpty by blast + next + case Empty + have "v3 = Stars []" by fact + then show "Stars (v # vs) :\val v3" + unfolding PosOrd_ex_eq_def using cond2 + by (simp add: PosOrd_shorterI) + qed +next + case (Posix_NTIMES2 vs r n v2) + then show "Stars vs :\val v2" + apply(simp add: LV_def) + apply(auto) + apply(erule Prf_elims) + apply(auto) + apply(rule PosOrd_eq_Stars_zipI) + prefer 2 + apply(simp) + prefer 2 + apply (metis Posix1(2) flats_empty) + apply(auto) + by (meson in_set_zipE) +next + case (Posix_UPNTIMES2 r n v2) + then show "Stars [] :\val v2" + apply(simp add: LV_def) + apply(auto) + apply(erule Prf_elims) + apply(auto) + unfolding PosOrd_ex_eq_def by simp +next + case (Posix_UPNTIMES1 s1 r v s2 n vs v3) + have "s1 \ r \ v" "s2 \ UPNTIMES r (n - 1) \ Stars vs" by fact+ + then have as1: "s1 = flat v" "s2 = flat (Stars vs)" by (auto dest: Posix1(2)) + have IH1: "\v3. v3 \ LV r s1 \ v :\val v3" by fact + have IH2: "\v3. v3 \ LV (UPNTIMES r (n - 1)) s2 \ Stars vs :\val v3" by fact + have cond: "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (UPNTIMES r (n - 1)))" by fact + have cond2: "flat v \ []" by fact + have "v3 \ LV (UPNTIMES r n) (s1 @ s2)" by fact + then consider + (NonEmpty) v3a vs3 where "v3 = Stars (v3a # vs3)" + "\ v3a : r" "\ Stars vs3 : UPNTIMES r (n - 1)" + "flats (v3a # vs3) = s1 @ s2" + | (Empty) "v3 = Stars []" + unfolding LV_def + apply(auto) + apply(erule Prf.cases) + apply(auto) + apply(case_tac vs) + apply(auto intro: Prf.intros) + by (simp add: Prf.intros(7) as1(1) cond2) + then show "Stars (v # vs) :\val v3" + proof (cases) + case (NonEmpty v3a vs3) + have "flats (v3a # vs3) = s1 @ s2" using NonEmpty(4) . + with cond have "flat v3a \pre s1" using NonEmpty(2,3) + unfolding prefix_list_def + apply(simp) + apply(simp add: append_eq_append_conv2) + apply(auto) + by (metis L_flat_Prf1 One_nat_def cond flat_Stars) + then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat (Stars vs3) = s2)" using NonEmpty(4) + by (simp add: sprefix_list_def append_eq_conv_conj) + then have q2: "v :\val v3a \ (flat v3a = s1 \ flat (Stars vs3) = s2)" + using PosOrd_spreI as1(1) NonEmpty(4) by blast + then have "v :\val v3a \ (v3a \ LV r s1 \ Stars vs3 \ LV (UPNTIMES r (n - 1)) s2)" + using NonEmpty(2,3) by (auto simp add: LV_def) + then have "v :\val v3a \ (v :\val v3a \ Stars vs :\val Stars vs3)" using IH1 IH2 by blast + then have "v :\val v3a \ (v = v3a \ Stars vs :\val Stars vs3)" + unfolding PosOrd_ex_eq_def by auto + then have "Stars (v # vs) :\val Stars (v3a # vs3)" using NonEmpty(4) q2 as1 + unfolding PosOrd_ex_eq_def + using PosOrd_StarsI PosOrd_StarsI2 by auto + then show "Stars (v # vs) :\val v3" unfolding NonEmpty by blast + next + case Empty + have "v3 = Stars []" by fact + then show "Stars (v # vs) :\val v3" + unfolding PosOrd_ex_eq_def using cond2 + by (simp add: PosOrd_shorterI) + qed +next + case (Posix_FROMNTIMES2 vs r n v2) + then show "Stars vs :\val v2" + apply(simp add: LV_def) + apply(auto) + apply(erule Prf_elims) + apply(auto) + apply(rule PosOrd_eq_Stars_zipI) + prefer 2 + apply(simp) + prefer 2 + apply (metis Posix1(2) flats_empty) + apply(auto) + by (meson in_set_zipE) +next + case (Posix_FROMNTIMES1 s1 r v s2 n vs v3) + have "s1 \ r \ v" "s2 \ FROMNTIMES r (n - 1) \ Stars vs" by fact+ + then have as1: "s1 = flat v" "s2 = flats vs" by (auto dest: Posix1(2)) + have IH1: "\v3. v3 \ LV r s1 \ v :\val v3" by fact + have IH2: "\v3. v3 \ LV (FROMNTIMES r (n - 1)) s2 \ Stars vs :\val v3" by fact + have cond: "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (FROMNTIMES r (n - 1)))" by fact + have cond2: "flat v \ []" by fact + have "v3 \ LV (FROMNTIMES r n) (s1 @ s2)" by fact + then consider + (NonEmpty) v3a vs3 where "v3 = Stars (v3a # vs3)" + "\ v3a : r" "\ Stars vs3 : FROMNTIMES r (n - 1)" + "flats (v3a # vs3) = s1 @ s2" + | (Empty) "v3 = Stars []" + unfolding LV_def + apply(auto) + apply(erule Prf.cases) + apply(auto) + apply(case_tac vs1) + apply(auto intro: Prf.intros) + apply(case_tac vs2) + apply(auto intro: Prf.intros) + apply (simp add: as1(1) cond2 flats_empty) + apply (simp add: Prf.intros) + apply(case_tac vs) + apply(auto) + using Posix_FROMNTIMES1.hyps(6) Prf.intros(10) by auto + then show "Stars (v # vs) :\val v3" + proof (cases) + case (NonEmpty v3a vs3) + have "flats (v3a # vs3) = s1 @ s2" using NonEmpty(4) . + with cond have "flat v3a \pre s1" using NonEmpty(2,3) + unfolding prefix_list_def + by (smt L_flat_Prf1 append_Nil2 append_eq_append_conv2 flat.simps(7) flat_Stars) + then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat (Stars vs3) = s2)" using NonEmpty(4) + by (simp add: sprefix_list_def append_eq_conv_conj) + then have q2: "v :\val v3a \ (flat v3a = s1 \ flat (Stars vs3) = s2)" + using PosOrd_spreI as1(1) NonEmpty(4) by blast + then have "v :\val v3a \ (v3a \ LV r s1 \ Stars vs3 \ LV (FROMNTIMES r (n - 1)) s2)" + using NonEmpty(2,3) by (auto simp add: LV_def) + then have "v :\val v3a \ (v :\val v3a \ Stars vs :\val Stars vs3)" using IH1 IH2 by blast + then have "v :\val v3a \ (v = v3a \ Stars vs :\val Stars vs3)" + unfolding PosOrd_ex_eq_def by auto + then have "Stars (v # vs) :\val Stars (v3a # vs3)" using NonEmpty(4) q2 as1 + unfolding PosOrd_ex_eq_def + using PosOrd_StarsI PosOrd_StarsI2 by auto + then show "Stars (v # vs) :\val v3" unfolding NonEmpty by blast + next + case Empty + have "v3 = Stars []" by fact + then show "Stars (v # vs) :\val v3" + unfolding PosOrd_ex_eq_def using cond2 + by (simp add: PosOrd_shorterI) + qed +next + case (Posix_NMTIMES2 vs r n m v2) + then show "Stars vs :\val v2" sorry +next + case (Posix_NMTIMES1 s1 r v s2 n m vs v2) + then show "Stars (v # vs) :\val v2" sorry +qed + + +lemma Posix_PosOrd_reverse: + assumes "s \ r \ v1" + shows "\(\v2 \ LV r s. v2 :\val v1)" +using assms +by (metis Posix_PosOrd less_irrefl PosOrd_def + PosOrd_ex_eq_def PosOrd_ex_def PosOrd_trans) + +lemma PosOrd_Posix: + assumes "v1 \ LV r s" "\v\<^sub>2 \ LV r s. \ v\<^sub>2 :\val v1" + shows "s \ r \ v1" +proof - + have "s \ L r" using assms(1) unfolding LV_def + using L_flat_Prf1 by blast + then obtain vposix where vp: "s \ r \ vposix" + using lexer_correct_Some by blast + with assms(1) have "vposix :\val v1" by (simp add: Posix_PosOrd) + then have "vposix = v1 \ vposix :\val v1" unfolding PosOrd_ex_eq2 by auto + moreover + { assume "vposix :\val v1" + moreover + have "vposix \ LV r s" using vp + using Posix_LV by blast + ultimately have "False" using assms(2) by blast + } + ultimately show "s \ r \ v1" using vp by blast +qed + +lemma Least_existence: + assumes "LV r s \ {}" + shows " \vmin \ LV r s. \v \ LV r s. vmin :\val v" +proof - + from assms + obtain vposix where "s \ r \ vposix" + unfolding LV_def + using L_flat_Prf1 lexer_correct_Some by blast + then have "\v \ LV r s. vposix :\val v" + by (simp add: Posix_PosOrd) + then show "\vmin \ LV r s. \v \ LV r s. vmin :\val v" + using Posix_LV \s \ r \ vposix\ by blast +qed + +lemma Least_existence1: + assumes "LV r s \ {}" + shows " \!vmin \ LV r s. \v \ LV r s. vmin :\val v" +using Least_existence[OF assms] assms + using PosOrdeq_antisym by blast + + + + + +lemma Least_existence1_pre: + assumes "LV r s \ {}" + shows " \!vmin \ LV r s. \v \ (LV r s \ {v'. flat v' \spre s}). vmin :\val v" +using Least_existence[OF assms] assms +apply - +apply(erule bexE) +apply(rule_tac a="vmin" in ex1I) +apply(auto)[1] +apply (metis PosOrd_Posix PosOrd_ex_eq2 PosOrd_spreI PosOrdeq_antisym Posix1(2)) +apply(auto)[1] +apply(simp add: PosOrdeq_antisym) +done + +lemma + shows "partial_order_on UNIV {(v1, v2). v1 :\val v2}" +apply(simp add: partial_order_on_def) +apply(simp add: preorder_on_def refl_on_def) +apply(simp add: PosOrdeq_refl) +apply(auto) +apply(rule transI) +apply(auto intro: PosOrdeq_trans)[1] +apply(rule antisymI) +apply(simp add: PosOrdeq_antisym) +done + +lemma + "wf {(v1, v2). v1 :\val v2 \ v1 \ LV r s \ v2 \ LV r s}" +apply(rule finite_acyclic_wf) +prefer 2 +apply(simp add: acyclic_def) +apply(induct_tac rule: trancl.induct) +apply(auto)[1] +oops + + +unused_thms + +end \ No newline at end of file diff -r deea42c83c9e -r a3134f7de065 thys/SpecExt.thy --- a/thys/SpecExt.thy Thu Oct 05 12:45:13 2017 +0100 +++ b/thys/SpecExt.thy Sat Oct 07 22:16:16 2017 +0100 @@ -311,8 +311,9 @@ apply(subst Pow_Sequ_Un) apply(simp) apply(simp only: Der_union Der_empty) -apply(simp) -apply(simp add: nullable_correctness del: Der_UNION) + apply(simp) +(* FROMNTIMES *) + apply(simp add: nullable_correctness del: Der_UNION) apply(subst Sequ_Union_in) apply(subst Der_Pow_Sequ[symmetric]) apply(subst Pow.simps[symmetric]) @@ -326,7 +327,8 @@ apply(auto simp add: Sequ_def)[1] apply(drule Pow_decomp) apply(auto)[1] -apply (metis append_Cons) + apply (metis append_Cons) +(* NMTIMES *) apply(simp add: nullable_correctness del: Der_UNION) apply(rule impI) apply(rule conjI) @@ -663,10 +665,10 @@ case (FROMNTIMES r n) have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact have "s \ L (FROMNTIMES r n)" by fact - then obtain ss1 ss2 m where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = m" "n \ m" + then obtain ss1 ss2 k where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = k" "n \ k" "\s \ set ss1. s \ L r \ s \ []" "\s \ set ss2. s \ L r \ s = []" using Pow_cstring by force - then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = m" "n \ m" + then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = k" "n \ k" "\v\set vs1. \ v : r \ flat v \ []" "\v\set vs2. \ v : r \ flat v = []" using IH flats_cval apply - @@ -723,11 +725,26 @@ apply(simp) done then show "\v. \ v : NMTIMES r n m \ flat v = s" - apply(rule_tac x="Stars (vs1 @ vs2)" in exI) + apply(case_tac "length vs1 \ n") + apply(rule_tac x="Stars (vs1 @ take (n - length vs1) vs2)" in exI) apply(simp) - apply(rule Prf.intros) - apply(auto) - sorry + apply(subgoal_tac "flats (take (n - length vs1) vs2) = []") + prefer 2 + apply (meson flats_empty in_set_takeD) + apply(clarify) + apply(rule conjI) + apply(rule Prf.intros) + apply(simp) + apply (meson in_set_takeD) + apply(simp) + apply(simp) + apply (simp add: flats_empty) + apply(rule_tac x="Stars vs1" in exI) + apply(simp) + apply(rule conjI) + apply(rule Prf.intros) + apply(auto) + done next case (UPNTIMES r n s) have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact @@ -806,6 +823,58 @@ ultimately show "finite (Prefixes s)" by simp qed +definition + "Stars_Cons V Vs \ {Stars (v # vs) | v vs. v \ V \ Stars vs \ Vs}" + +definition + "Stars_Append Vs1 Vs2 \ {Stars (vs1 @ vs2) | vs1 vs2. Stars vs1 \ Vs1 \ Stars vs2 \ Vs2}" + +fun Stars_Pow :: "val set \ nat \ val set" +where + "Stars_Pow Vs 0 = {Stars []}" +| "Stars_Pow Vs (Suc n) = Stars_Cons Vs (Stars_Pow Vs n)" + +lemma finite_Stars_Cons: + assumes "finite V" "finite Vs" + shows "finite (Stars_Cons V Vs)" + using assms +proof - + from assms(2) have "finite (Stars -` Vs)" + by(simp add: finite_vimageI inj_on_def) + with assms(1) have "finite (V \ (Stars -` Vs))" + by(simp) + then have "finite ((\(v, vs). Stars (v # vs)) ` (V \ (Stars -` Vs)))" + by simp + moreover have "Stars_Cons V Vs = (\(v, vs). Stars (v # vs)) ` (V \ (Stars -` Vs))" + unfolding Stars_Cons_def by auto + ultimately show "finite (Stars_Cons V Vs)" + by simp +qed + +lemma finite_Stars_Append: + assumes "finite Vs1" "finite Vs2" + shows "finite (Stars_Append Vs1 Vs2)" + using assms +proof - + define UVs1 where "UVs1 \ Stars -` Vs1" + define UVs2 where "UVs2 \ Stars -` Vs2" + from assms have "finite UVs1" "finite UVs2" + unfolding UVs1_def UVs2_def + by(simp_all add: finite_vimageI inj_on_def) + then have "finite ((\(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \ UVs2))" + by simp + moreover + have "Stars_Append Vs1 Vs2 = (\(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \ UVs2)" + unfolding Stars_Append_def UVs1_def UVs2_def by auto + ultimately show "finite (Stars_Append Vs1 Vs2)" + by simp +qed + +lemma finite_Stars_Pow: + assumes "finite Vs" + shows "finite (Stars_Pow Vs n)" +by (induct n) (simp_all add: finite_Stars_Cons assms) + lemma LV_STAR_finite: assumes "\s. finite (LV r s)" shows "finite (LV (STAR r) s)" @@ -816,18 +885,20 @@ by (auto simp add: strict_suffix_def) define f where "f \ \(v, vs). Stars (v # vs)" define S1 where "S1 \ \s' \ Prefixes s. LV r s'" - define S2 where "S2 \ \s2 \ SSuffixes s. Stars -` (LV (STAR r) s2)" + define S2 where "S2 \ \s2 \ SSuffixes s. LV (STAR r) s2" have "finite S1" using assms unfolding S1_def by (simp_all add: finite_Prefixes) moreover with IH have "finite S2" unfolding S2_def - by (auto simp add: finite_SSuffixes inj_on_def finite_vimageI) + by (auto simp add: finite_SSuffixes) ultimately - have "finite ({Stars []} \ f ` (S1 \ S2))" by simp + have "finite ({Stars []} \ Stars_Cons S1 S2)" + by (simp add: finite_Stars_Cons) moreover - have "LV (STAR r) s \ {Stars []} \ f ` (S1 \ S2)" - unfolding S1_def S2_def f_def - unfolding LV_def image_def prefix_def strict_suffix_def + have "LV (STAR r) s \ {Stars []} \ (Stars_Cons S1 S2)" + unfolding S1_def S2_def f_def LV_def Stars_Cons_def + unfolding prefix_def strict_suffix_def + unfolding image_def apply(auto) apply(case_tac x) apply(auto elim: Prf_elims) @@ -837,12 +908,12 @@ apply(auto intro: Prf.intros) apply(rule exI) apply(rule conjI) - apply(rule_tac x="flat a" in exI) + apply(rule_tac x="flats list" in exI) apply(rule conjI) - apply(rule_tac x="flats list" in exI) + apply(rule_tac x="flat a" in exI) apply(simp) apply(blast) - using Prf.intros(6) by blast + using Prf.intros(6) flat_Stars by blast ultimately show "finite (LV (STAR r) s)" by (simp add: finite_subset) qed @@ -851,37 +922,6 @@ "LV (UPNTIMES r n) s \ LV (STAR r) s" by(auto simp add: LV_def intro: Prf.intros elim: Prf_elims) - -lemma LV_NTIMES_0: - shows "LV (NTIMES r 0) s \ {Stars []}" -unfolding LV_def -apply(auto elim: Prf_elims) -done - -lemma LV_NTIMES_1: - shows "LV (NTIMES r 1) s \ (\v. Stars [v]) ` (LV r s)" -unfolding LV_def -apply(auto elim!: Prf_elims) -apply(case_tac vs1) -apply(simp) -apply(case_tac vs2) -apply(simp) -apply(simp) -apply(simp) -done - -lemma LV_NTIMES_2: - shows "LV (NTIMES r 2) [] \ (\(v1,v2). Stars [v1,v2]) ` (LV r [] \ LV r [])" -unfolding LV_def -apply(auto elim!: Prf_elims simp add: image_def) -apply(case_tac vs1) -apply(auto) -apply(case_tac vs2) -apply(auto) -apply(case_tac list) -apply(auto) -done - lemma LV_NTIMES_3: shows "LV (NTIMES r (Suc n)) [] = (\(v,vs). Stars (v#vs)) ` (LV r [] \ (Stars -` (LV (NTIMES r n) [])))" unfolding LV_def @@ -896,66 +936,147 @@ apply(subst append.simps(1)[symmetric]) apply(rule Prf.intros) apply(auto) -done - -thm card_cartesian_product - -lemma finite_list: - assumes "finite A" - shows "finite {vs. \v\set vs. v \ A \ length vs = n}" - apply(induct n) - apply(simp) - apply (smt Collect_cong empty_iff finite.emptyI finite.insertI - in_listsI list.set(1) lists_empty mem_Collect_eq singleton_conv2) - apply(rule_tac B="{[]} \ (\(v,vs). v # vs) `(A \ {vs. \v\set vs. v \ A \ length vs = n})" in finite_subset) - apply(auto simp add: image_def)[1] - apply(case_tac x) - apply(simp) - apply(simp) - apply(simp) - apply(rule finite_imageI) - using assms - apply(simp) - done - -lemma test: - "LV (NTIMES r n) [] \ Stars ` {vs. \v \ set vs. v \ LV r [] \ length vs = n}" -apply(auto simp add: LV_def elim: Prf_elims) -done + done -lemma test3: - "LV (FROMNTIMES r n) [] \ Stars ` {vs. \v \ set vs. v \ LV r [] \ length vs = n}" - apply(auto simp add: image_def LV_def elim!: Prf_elims) - apply blast - apply(case_tac vs) +lemma LV_FROMNTIMES_3: + shows "LV (FROMNTIMES r (Suc n)) [] = + (\(v,vs). Stars (v#vs)) ` (LV r [] \ (Stars -` (LV (FROMNTIMES r n) [])))" +unfolding LV_def +apply(auto elim!: Prf_elims simp add: image_def) +apply(case_tac vs1) +apply(auto) +apply(case_tac vs2) +apply(auto) +apply(subst append.simps(1)[symmetric]) +apply(rule Prf.intros) + apply(auto) + apply (metis le_imp_less_Suc length_greater_0_conv less_antisym list.exhaust list.set_intros(1) not_less_eq zero_le) + prefer 2 + using nth_mem apply blast + apply(case_tac vs1) + 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)) apply(auto) - done - - +done + +lemma LV_NTIMES_4: + "LV (NTIMES r n) [] = Stars_Pow (LV r []) n" + apply(induct n) + apply(simp add: LV_def) + apply(auto elim!: Prf_elims simp add: image_def)[1] + apply(subst append.simps[symmetric]) + apply(rule Prf.intros) + apply(simp_all) + apply(simp add: LV_NTIMES_3 image_def Stars_Cons_def) + apply blast + done -lemma LV_NTIMES_empty_finite: - assumes "finite (LV r [])" - shows "finite (LV (NTIMES r n) [])" - using assms - apply - +lemma LV_NTIMES_5: + "LV (NTIMES r n) s \ Stars_Append (LV (STAR r) s) (\i\n. LV (NTIMES r i) [])" +apply(auto simp add: LV_def) +apply(auto elim!: Prf_elims) + apply(auto simp add: Stars_Append_def) + apply(rule_tac x="vs1" in exI) + apply(rule_tac x="vs2" in exI) + apply(auto) + using Prf.intros(6) apply(auto) + apply(rule_tac x="length vs2" in bexI) + thm Prf.intros + apply(subst append.simps(1)[symmetric]) + apply(rule Prf.intros) + apply(auto)[1] + apply(auto)[1] + apply(simp) + apply(simp) + done + +lemma ttty: + "LV (FROMNTIMES r n) [] = Stars_Pow (LV r []) n" + apply(induct n) + apply(simp add: LV_def) + apply(auto elim: Prf_elims simp add: image_def)[1] + prefer 2 + apply(subst append.simps[symmetric]) + apply(rule Prf.intros) + apply(simp_all) + apply(erule Prf_elims) + apply(case_tac vs1) + apply(simp) + apply(simp) + apply(case_tac x) + apply(simp_all) + apply(simp add: LV_FROMNTIMES_3 image_def Stars_Cons_def) + apply blast + done + +lemma LV_FROMNTIMES_5: + "LV (FROMNTIMES r n) s \ Stars_Append (LV (STAR r) s) (\i\n. LV (FROMNTIMES r i) [])" +apply(auto simp add: LV_def) +apply(auto elim!: Prf_elims) + apply(auto simp add: Stars_Append_def) + apply(rule_tac x="vs1" in exI) + apply(rule_tac x="vs2" in exI) + apply(auto) + using Prf.intros(6) apply(auto) + apply(rule_tac x="length vs2" in bexI) + thm Prf.intros + apply(subst append.simps(1)[symmetric]) + apply(rule Prf.intros) + apply(auto)[1] + apply(auto)[1] + apply(simp) + apply(simp) + apply(rule_tac x="vs" in exI) + apply(rule_tac x="[]" in exI) + apply(auto) + by (metis Prf.intros(9) append_Nil atMost_iff empty_iff le_imp_less_Suc less_antisym list.set(1) nth_mem zero_le) + +lemma LV_FROMNTIMES_6: + assumes "\s. finite (LV r s)" + shows "finite (LV (FROMNTIMES r n) s)" apply(rule finite_subset) - apply(rule test) - apply(rule finite_imageI) - apply(rule finite_list) - apply(simp) - done + apply(rule LV_FROMNTIMES_5) + apply(rule finite_Stars_Append) + apply(rule LV_STAR_finite) + apply(rule assms) + apply(rule finite_UN_I) + apply(auto) + by (simp add: assms finite_Stars_Pow ttty) -lemma LV_NTIMES_STAR: - "LV (NTIMES r n) s \ LV (STAR r) s" -apply(auto simp add: LV_def intro: Prf.intros elim!: Prf_elims) -apply(rule Prf.intros) -oops +lemma LV_NMTIMES_5: + "LV (NMTIMES r n m) s \ Stars_Append (LV (STAR r) s) (\i\n. LV (FROMNTIMES r i) [])" +apply(auto simp add: LV_def) +apply(auto elim!: Prf_elims) + apply(auto simp add: Stars_Append_def) + apply(rule_tac x="vs1" in exI) + apply(rule_tac x="vs2" in exI) + apply(auto) + using Prf.intros(6) apply(auto) + apply(rule_tac x="length vs2" in bexI) + thm Prf.intros + apply(subst append.simps(1)[symmetric]) + apply(rule Prf.intros) + apply(auto)[1] + apply(auto)[1] + apply(simp) + apply(simp) + apply(rule_tac x="vs" in exI) + apply(rule_tac x="[]" in exI) + apply(auto) + by (metis Prf.intros(9) append_Nil atMost_iff empty_iff le_imp_less_Suc less_antisym list.set(1) nth_mem zero_le) -lemma LV_FROMNTIMES_STAR: - "LV (FROMNTIMES r n) s \ LV (STAR r) s" -apply(auto simp add: LV_def intro: Prf.intros elim!: Prf_elims) -oops - +lemma LV_NMTIMES_6: + assumes "\s. finite (LV r s)" + shows "finite (LV (NMTIMES r n m) s)" + apply(rule finite_subset) + apply(rule LV_NMTIMES_5) + apply(rule finite_Stars_Append) + apply(rule LV_STAR_finite) + apply(rule assms) + apply(rule finite_UN_I) + apply(auto) + by (simp add: assms finite_Stars_Pow ttty) + + lemma LV_finite: shows "finite (LV r s)" proof(induct r arbitrary: s) @@ -999,7 +1120,17 @@ case (FROMNTIMES r n s) have "\s. finite (LV r s)" by fact then show "finite (LV (FROMNTIMES r n) s)" - + by (simp add: LV_FROMNTIMES_6) +next + case (NTIMES r n s) + have "\s. finite (LV r s)" by fact + then show "finite (LV (NTIMES r n) s)" + 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) +next + case (NMTIMES r n m s) + have "\s. finite (LV r s)" by fact + then show "finite (LV (NMTIMES r n m) s)" + by (simp add: LV_NMTIMES_6) qed @@ -1020,7 +1151,26 @@ \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (STAR r))\ \ (s1 @ s2) \ STAR r \ Stars (v # vs)" | Posix_STAR2: "[] \ STAR r \ Stars []" - +| Posix_NTIMES1: "\s1 \ r \ v; s2 \ NTIMES r (n - 1) \ Stars vs; flat v \ []; 0 < n; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (NTIMES r (n - 1)))\ + \ (s1 @ s2) \ NTIMES r n \ Stars (v # vs)" +| Posix_NTIMES2: "\\v \ set vs. [] \ r \ v; length vs = n\ + \ [] \ NTIMES r n \ Stars vs" +| Posix_UPNTIMES1: "\s1 \ r \ v; s2 \ UPNTIMES r (n - 1) \ Stars vs; flat v \ []; 0 < n; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (UPNTIMES r (n - 1)))\ + \ (s1 @ s2) \ UPNTIMES r n \ Stars (v # vs)" +| Posix_UPNTIMES2: "[] \ UPNTIMES r n \ Stars []" +| Posix_FROMNTIMES2: "\\v \ set vs. [] \ r \ v; length vs = n\ + \ [] \ FROMNTIMES r n \ Stars vs" +| Posix_FROMNTIMES1: "\s1 \ r \ v; s2 \ FROMNTIMES r (n - 1) \ Stars vs; flat v \ []; 0 < n; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (FROMNTIMES r (n - 1)))\ + \ (s1 @ s2) \ FROMNTIMES r n \ Stars (v # vs)" +| Posix_NMTIMES2: "\\v \ set vs. [] \ r \ v; length vs = n; n \ m\ + \ [] \ NMTIMES r n m \ Stars vs" +| Posix_NMTIMES1: "\s1 \ r \ v; s2 \ NMTIMES r n m \ Stars vs; flat v \ []; n \ m; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (NMTIMES r n m))\ + \ (s1 @ s2) \ NMTIMES r (Suc n) (Suc m) \ Stars (v # vs)" + inductive_cases Posix_elims: "s \ ZERO \ v" "s \ ONE \ v" @@ -1028,17 +1178,60 @@ "s \ ALT r1 r2 \ v" "s \ SEQ r1 r2 \ v" "s \ STAR r \ v" - + "s \ NTIMES r n \ v" + "s \ UPNTIMES r n \ v" + "s \ FROMNTIMES r n \ v" + "s \ NMTIMES r n m \ v" + lemma Posix1: 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) - + apply(induct s r v rule: Posix.induct) + apply(auto simp add: Sequ_def)[18] + apply(case_tac n) + apply(simp) + apply(simp add: Sequ_def) + apply(auto)[1] + apply(simp) + apply(clarify) + apply(rule_tac x="Suc x" in bexI) + apply(simp add: Sequ_def) + apply(auto)[5] + using nth_mem nullable.simps(9) nullable_correctness apply auto[1] + apply simp + apply(simp) + apply(clarify) + apply(rule_tac x="Suc x" in bexI) + apply(simp add: Sequ_def) + apply(auto)[3] + apply(simp) + apply fastforce + apply(simp) + apply(simp) + apply(clarify) + apply(rule_tac x="Suc x" in bexI) + apply(auto simp add: Sequ_def)[2] + apply(simp) +done + text {* Our Posix definition determines a unique value. *} + +lemma List_eq_zipI: + assumes "\(v1, v2) \ set (zip vs1 vs2). v1 = v2" + and "length vs1 = length vs2" + shows "vs1 = vs2" + using assms + apply(induct vs1 arbitrary: vs2) + apply(case_tac vs2) + apply(simp) + apply(simp) + apply(case_tac vs2) + apply(simp) + apply(simp) +done lemma Posix_determ: assumes "s \ r \ v1" "s \ r \ v2" @@ -1104,6 +1297,89 @@ case (Posix_STAR2 r v2) have "[] \ STAR r \ v2" by fact then show "Stars [] = v2" by cases (auto simp add: Posix1) +next + case (Posix_NTIMES2 vs r n v2) + then show "Stars vs = v2" + apply(erule_tac Posix_elims) + apply(auto) + apply (simp add: Posix1(2)) + apply(rule List_eq_zipI) + apply(auto) + by (meson in_set_zipE) +next + case (Posix_NTIMES1 s1 r v s2 n vs v2) + have "(s1 @ s2) \ NTIMES r n \ v2" + "s1 \ r \ v" "s2 \ NTIMES r (n - 1) \ Stars vs" "flat v \ []" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (NTIMES r (n - 1 )))" by fact+ + then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (NTIMES r (n - 1)) \ (Stars vs')" + apply(cases) apply (auto simp add: append_eq_append_conv2) + using Posix1(1) apply fastforce + apply (metis One_nat_def Posix1(1) Posix_NTIMES1.hyps(7) append.right_neutral append_self_conv2) + using Posix1(2) by blast + moreover + have IHs: "\v2. s1 \ r \ v2 \ v = v2" + "\v2. s2 \ NTIMES r (n - 1) \ v2 \ Stars vs = v2" by fact+ + ultimately show "Stars (v # vs) = v2" by auto +next + case (Posix_UPNTIMES1 s1 r v s2 n vs v2) + have "(s1 @ s2) \ UPNTIMES r n \ v2" + "s1 \ r \ v" "s2 \ UPNTIMES r (n - 1) \ Stars vs" "flat v \ []" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (UPNTIMES r (n - 1 )))" by fact+ + then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (UPNTIMES r (n - 1)) \ (Stars vs')" + apply(cases) apply (auto simp add: append_eq_append_conv2) + using Posix1(1) apply fastforce + apply (metis One_nat_def Posix1(1) Posix_UPNTIMES1.hyps(7) append.right_neutral append_self_conv2) + using Posix1(2) by blast + moreover + have IHs: "\v2. s1 \ r \ v2 \ v = v2" + "\v2. s2 \ UPNTIMES r (n - 1) \ v2 \ Stars vs = v2" by fact+ + ultimately show "Stars (v # vs) = v2" by auto +next + case (Posix_UPNTIMES2 r n v2) + then show "Stars [] = v2" + apply(erule_tac Posix_elims) + apply(auto) + by (simp add: Posix1(2)) +next + case (Posix_FROMNTIMES1 s1 r v s2 n vs v2) + have "(s1 @ s2) \ FROMNTIMES r n \ v2" + "s1 \ r \ v" "s2 \ FROMNTIMES r (n - 1) \ Stars vs" "flat v \ []" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (FROMNTIMES r (n - 1 )))" by fact+ + then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (FROMNTIMES r (n - 1)) \ (Stars vs')" + apply(cases) apply (auto simp add: append_eq_append_conv2) + using Posix1(1) Posix1(2) apply blast + apply(drule_tac x="v" in meta_spec) + apply(drule_tac x="vs" in meta_spec) + apply(simp) + apply(drule meta_mp) + apply(case_tac n) + apply(simp) + apply(rule conjI) + apply (smt L.simps(9) One_nat_def Posix1(1) Posix_FROMNTIMES1.hyps Suc_mono Suc_pred UN_E append.right_neutral append_Nil lessI less_antisym list.size(3) nat.simps(3) neq0_conv not_less_eq val.inject(5)) + apply(rule List_eq_zipI) + apply(auto)[1] + sorry + moreover + have IHs: "\v2. s1 \ r \ v2 \ v = v2" + "\v2. s2 \ FROMNTIMES r (n - 1) \ v2 \ Stars vs = v2" by fact+ + ultimately show "Stars (v # vs) = v2" by auto +next + case (Posix_FROMNTIMES2 vs r n v2) + then show "Stars vs = v2" + apply(erule_tac Posix_elims) + apply(auto) + apply(rule List_eq_zipI) + apply(auto) + apply(meson in_set_zipE) + by (simp add: Posix1(2)) +next + case (Posix_NMTIMES1 s1 r v s2 n m vs v2) + then show "Stars (v # vs) = v2" + sorry +next + case (Posix_NMTIMES2 vs r n m v2) + then show "Stars vs = v2" + sorry qed @@ -1116,7 +1392,60 @@ shows "v \ LV r s" using assms unfolding LV_def apply(induct rule: Posix.induct) -apply(auto simp add: intro!: Prf.intros elim!: Prf_elims) -done - + apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)[7] + defer + defer + apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)[2] + apply (metis (mono_tags, lifting) Prf.intros(9) append_Nil empty_iff flat_Stars flats_empty list.set(1) mem_Collect_eq) + prefer 4 + apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)[1] + apply(subst append.simps(2)[symmetric]) + apply(rule Prf.intros) + apply(auto) + prefer 4 + apply (metis Prf.intros(8) append_Nil empty_iff list.set(1)) + apply(erule Posix_elims) + apply(auto) + apply(rule_tac t="v # vsa" and s = "[v] @ vsa" in subst) + apply(simp) + apply(rule Prf.intros) + apply(simp) + apply(auto)[1] + apply(auto simp add: Sequ_def intro: Prf.intros elim: Prf_elims)[1] + apply(simp) + apply(rotate_tac 4) + apply(erule Prf_elims) + apply(auto) + apply(case_tac n) + apply(simp) + + apply(subst append.simps(2)[symmetric]) + apply(rule Prf.intros) + apply(auto) + apply(rule Prf.intros(10)) + apply(auto) + apply (metis Prf.intros(11) append_Nil empty_iff list.set(1)) + apply(rotate_tac 6) + apply(erule Prf_elims) + apply(simp) + apply(subst append.simps(2)[symmetric]) + apply(rule Prf.intros) + apply(auto) + apply(rule_tac t="v # vsa" and s = "(v # vsa) @ []" in subst) + apply(simp) + apply(simp add: Prf.intros(12)) + done + +lemma FROMNTIMES_0: + assumes "s \ FROMNTIMES r 0 \ v" + shows "s = [] \ v = Stars []" + using assms + apply(erule_tac Posix_elims) + apply(auto) + done + +lemma FROMNTIMES_der_0: + shows "der c (FROMNTIMES r 0) = SEQ (der c r) (FROMNTIMES r 0)" +by(simp) + end \ No newline at end of file diff -r deea42c83c9e -r a3134f7de065 thys/journal/llncs.cls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/journal/llncs.cls Sat Oct 07 22:16:16 2017 +0100 @@ -0,0 +1,1189 @@ +% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002) +% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science +% +%% +%% \CharacterTable +%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z +%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z +%% Digits \0\1\2\3\4\5\6\7\8\9 +%% Exclamation \! Double quote \" Hash (number) \# +%% Dollar \$ Percent \% Ampersand \& +%% Acute accent \' Left paren \( Right paren \) +%% Asterisk \* Plus \+ Comma \, +%% Minus \- Point \. Solidus \/ +%% Colon \: Semicolon \; Less than \< +%% Equals \= Greater than \> Question mark \? +%% Commercial at \@ Left bracket \[ Backslash \\ +%% Right bracket \] Circumflex \^ Underscore \_ +%% Grave accent \` Left brace \{ Vertical bar \| +%% Right brace \} Tilde \~} +%% +\NeedsTeXFormat{LaTeX2e}[1995/12/01] +\ProvidesClass{llncs}[2002/01/28 v2.13 +^^J LaTeX document class for Lecture Notes in Computer Science] +% Options +\let\if@envcntreset\iffalse +\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue} +\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y} +\DeclareOption{oribibl}{\let\oribibl=Y} +\let\if@custvec\iftrue +\DeclareOption{orivec}{\let\if@custvec\iffalse} +\let\if@envcntsame\iffalse +\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue} +\let\if@envcntsect\iffalse +\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue} +\let\if@runhead\iffalse +\DeclareOption{runningheads}{\let\if@runhead\iftrue} + +\let\if@openbib\iffalse +\DeclareOption{openbib}{\let\if@openbib\iftrue} + +% languages +\let\switcht@@therlang\relax +\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}} +\def\ds@francais{\def\switcht@@therlang{\switcht@francais}} + +\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}} + +\ProcessOptions + +\LoadClass[twoside]{article} +\RequirePackage{multicol} % needed for the list of participants, index + +\setlength{\textwidth}{12.2cm} +\setlength{\textheight}{19.3cm} +\renewcommand\@pnumwidth{2em} +\renewcommand\@tocrmarg{3.5em} +% +\def\@dottedtocline#1#2#3#4#5{% + \ifnum #1>\c@tocdepth \else + \vskip \z@ \@plus.2\p@ + {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm + \parfillskip -\rightskip \pretolerance=10000 + \parindent #2\relax\@afterindenttrue + \interlinepenalty\@M + \leavevmode + \@tempdima #3\relax + \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip + {#4}\nobreak + \leaders\hbox{$\m@th + \mkern \@dotsep mu\hbox{.}\mkern \@dotsep + mu$}\hfill + \nobreak + \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}% + \par}% + \fi} +% +\def\switcht@albion{% +\def\abstractname{Abstract.} +\def\ackname{Acknowledgement.} +\def\andname{and} +\def\lastandname{\unskip, and} +\def\appendixname{Appendix} +\def\chaptername{Chapter} +\def\claimname{Claim} +\def\conjecturename{Conjecture} +\def\contentsname{Table of Contents} +\def\corollaryname{Corollary} +\def\definitionname{Definition} +\def\examplename{Example} +\def\exercisename{Exercise} +\def\figurename{Fig.} +\def\keywordname{{\bf Key words:}} +\def\indexname{Index} +\def\lemmaname{Lemma} +\def\contriblistname{List of Contributors} +\def\listfigurename{List of Figures} +\def\listtablename{List of Tables} +\def\mailname{{\it Correspondence to\/}:} +\def\noteaddname{Note added in proof} +\def\notename{Note} +\def\partname{Part} +\def\problemname{Problem} +\def\proofname{Proof} +\def\propertyname{Property} +\def\propositionname{Proposition} +\def\questionname{Question} +\def\remarkname{Remark} +\def\seename{see} +\def\solutionname{Solution} +\def\subclassname{{\it Subject Classifications\/}:} +\def\tablename{Table} +\def\theoremname{Theorem}} +\switcht@albion +% Names of theorem like environments are already defined +% but must be translated if another language is chosen +% +% French section +\def\switcht@francais{%\typeout{On parle francais.}% + \def\abstractname{R\'esum\'e.}% + \def\ackname{Remerciements.}% + \def\andname{et}% + \def\lastandname{ et}% + \def\appendixname{Appendice} + \def\chaptername{Chapitre}% + \def\claimname{Pr\'etention}% + \def\conjecturename{Hypoth\`ese}% + \def\contentsname{Table des mati\`eres}% + \def\corollaryname{Corollaire}% + \def\definitionname{D\'efinition}% + \def\examplename{Exemple}% + \def\exercisename{Exercice}% + \def\figurename{Fig.}% + \def\keywordname{{\bf Mots-cl\'e:}} + \def\indexname{Index} + \def\lemmaname{Lemme}% + \def\contriblistname{Liste des contributeurs} + \def\listfigurename{Liste des figures}% + \def\listtablename{Liste des tables}% + \def\mailname{{\it Correspondence to\/}:} + \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}% + \def\notename{Remarque}% + \def\partname{Partie}% + \def\problemname{Probl\`eme}% + \def\proofname{Preuve}% + \def\propertyname{Caract\'eristique}% +%\def\propositionname{Proposition}% + \def\questionname{Question}% + \def\remarkname{Remarque}% + \def\seename{voir} + \def\solutionname{Solution}% + \def\subclassname{{\it Subject Classifications\/}:} + \def\tablename{Tableau}% + \def\theoremname{Th\'eor\`eme}% +} +% +% German section +\def\switcht@deutsch{%\typeout{Man spricht deutsch.}% + \def\abstractname{Zusammenfassung.}% + \def\ackname{Danksagung.}% + \def\andname{und}% + \def\lastandname{ und}% + \def\appendixname{Anhang}% + \def\chaptername{Kapitel}% + \def\claimname{Behauptung}% + \def\conjecturename{Hypothese}% + \def\contentsname{Inhaltsverzeichnis}% + \def\corollaryname{Korollar}% +%\def\definitionname{Definition}% + \def\examplename{Beispiel}% + \def\exercisename{\"Ubung}% + \def\figurename{Abb.}% + \def\keywordname{{\bf Schl\"usselw\"orter:}} + \def\indexname{Index} +%\def\lemmaname{Lemma}% + \def\contriblistname{Mitarbeiter} + \def\listfigurename{Abbildungsverzeichnis}% + \def\listtablename{Tabellenverzeichnis}% + \def\mailname{{\it Correspondence to\/}:} + \def\noteaddname{Nachtrag}% + \def\notename{Anmerkung}% + \def\partname{Teil}% +%\def\problemname{Problem}% + \def\proofname{Beweis}% + \def\propertyname{Eigenschaft}% +%\def\propositionname{Proposition}% + \def\questionname{Frage}% + \def\remarkname{Anmerkung}% + \def\seename{siehe} + \def\solutionname{L\"osung}% + \def\subclassname{{\it Subject Classifications\/}:} + \def\tablename{Tabelle}% +%\def\theoremname{Theorem}% +} + +% Ragged bottom for the actual page +\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil +\global\let\@textbottom\relax}} + +\renewcommand\small{% + \@setfontsize\small\@ixpt{11}% + \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@ + \abovedisplayshortskip \z@ \@plus2\p@ + \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@ + \def\@listi{\leftmargin\leftmargini + \parsep 0\p@ \@plus1\p@ \@minus\p@ + \topsep 8\p@ \@plus2\p@ \@minus4\p@ + \itemsep0\p@}% + \belowdisplayskip \abovedisplayskip +} + +\frenchspacing +\widowpenalty=10000 +\clubpenalty=10000 + +\setlength\oddsidemargin {63\p@} +\setlength\evensidemargin {63\p@} +\setlength\marginparwidth {90\p@} + +\setlength\headsep {16\p@} + +\setlength\footnotesep{7.7\p@} +\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@} +\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@} + +\setcounter{secnumdepth}{2} + +\newcounter {chapter} +\renewcommand\thechapter {\@arabic\c@chapter} + +\newif\if@mainmatter \@mainmattertrue +\newcommand\frontmatter{\cleardoublepage + \@mainmatterfalse\pagenumbering{Roman}} +\newcommand\mainmatter{\cleardoublepage + \@mainmattertrue\pagenumbering{arabic}} +\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi + \@mainmatterfalse} + +\renewcommand\part{\cleardoublepage + \thispagestyle{empty}% + \if@twocolumn + \onecolumn + \@tempswatrue + \else + \@tempswafalse + \fi + \null\vfil + \secdef\@part\@spart} + +\def\@part[#1]#2{% + \ifnum \c@secnumdepth >-2\relax + \refstepcounter{part}% + \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}% + \else + \addcontentsline{toc}{part}{#1}% + \fi + \markboth{}{}% + {\centering + \interlinepenalty \@M + \normalfont + \ifnum \c@secnumdepth >-2\relax + \huge\bfseries \partname~\thepart + \par + \vskip 20\p@ + \fi + \Huge \bfseries #2\par}% + \@endpart} +\def\@spart#1{% + {\centering + \interlinepenalty \@M + \normalfont + \Huge \bfseries #1\par}% + \@endpart} +\def\@endpart{\vfil\newpage + \if@twoside + \null + \thispagestyle{empty}% + \newpage + \fi + \if@tempswa + \twocolumn + \fi} + +\newcommand\chapter{\clearpage + \thispagestyle{empty}% + \global\@topnum\z@ + \@afterindentfalse + \secdef\@chapter\@schapter} +\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne + \if@mainmatter + \refstepcounter{chapter}% + \typeout{\@chapapp\space\thechapter.}% + \addcontentsline{toc}{chapter}% + {\protect\numberline{\thechapter}#1}% + \else + \addcontentsline{toc}{chapter}{#1}% + \fi + \else + \addcontentsline{toc}{chapter}{#1}% + \fi + \chaptermark{#1}% + \addtocontents{lof}{\protect\addvspace{10\p@}}% + \addtocontents{lot}{\protect\addvspace{10\p@}}% + \if@twocolumn + \@topnewpage[\@makechapterhead{#2}]% + \else + \@makechapterhead{#2}% + \@afterheading + \fi} +\def\@makechapterhead#1{% +% \vspace*{50\p@}% + {\centering + \ifnum \c@secnumdepth >\m@ne + \if@mainmatter + \large\bfseries \@chapapp{} \thechapter + \par\nobreak + \vskip 20\p@ + \fi + \fi + \interlinepenalty\@M + \Large \bfseries #1\par\nobreak + \vskip 40\p@ + }} +\def\@schapter#1{\if@twocolumn + \@topnewpage[\@makeschapterhead{#1}]% + \else + \@makeschapterhead{#1}% + \@afterheading + \fi} +\def\@makeschapterhead#1{% +% \vspace*{50\p@}% + {\centering + \normalfont + \interlinepenalty\@M + \Large \bfseries #1\par\nobreak + \vskip 40\p@ + }} + +\renewcommand\section{\@startsection{section}{1}{\z@}% + {-18\p@ \@plus -4\p@ \@minus -4\p@}% + {12\p@ \@plus 4\p@ \@minus 4\p@}% + {\normalfont\large\bfseries\boldmath + \rightskip=\z@ \@plus 8em\pretolerance=10000 }} +\renewcommand\subsection{\@startsection{subsection}{2}{\z@}% + {-18\p@ \@plus -4\p@ \@minus -4\p@}% + {8\p@ \@plus 4\p@ \@minus 4\p@}% + {\normalfont\normalsize\bfseries\boldmath + \rightskip=\z@ \@plus 8em\pretolerance=10000 }} +\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}% + {-18\p@ \@plus -4\p@ \@minus -4\p@}% + {-0.5em \@plus -0.22em \@minus -0.1em}% + {\normalfont\normalsize\bfseries\boldmath}} +\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}% + {-12\p@ \@plus -4\p@ \@minus -4\p@}% + {-0.5em \@plus -0.22em \@minus -0.1em}% + {\normalfont\normalsize\itshape}} +\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use + \string\subparagraph\space with this class}\vskip0.5cm +You should not use \verb|\subparagraph| with this class.\vskip0.5cm} + +\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00} +\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01} +\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02} +\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03} +\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04} +\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05} +\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06} +\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07} +\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08} +\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09} +\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A} + +\let\footnotesize\small + +\if@custvec +\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}} +{\mbox{\boldmath$\textstyle#1$}} +{\mbox{\boldmath$\scriptstyle#1$}} +{\mbox{\boldmath$\scriptscriptstyle#1$}}} +\fi + +\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}} +\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil +\penalty50\hskip1em\null\nobreak\hfil\squareforqed +\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi} + +\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip +\halign{\hfil +$\displaystyle##$\hfil\cr\gets\cr\to\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets +\cr\to\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets +\cr\to\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +\gets\cr\to\cr}}}}} +\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil +$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr +\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr +\noalign{\vskip1pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +<\cr +\noalign{\vskip0.9pt}=\cr}}}}} +\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil +$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr +\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr +\noalign{\vskip1pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +>\cr +\noalign{\vskip0.9pt}=\cr}}}}} +\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip +\halign{\hfil +$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr +>\cr\noalign{\vskip-1pt}<\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr +>\cr\noalign{\vskip-0.8pt}<\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +>\cr\noalign{\vskip-0.3pt}<\cr}}}}} +\def\bbbr{{\rm I\!R}} %reelle Zahlen +\def\bbbm{{\rm I\!M}} +\def\bbbn{{\rm I\!N}} %natuerliche Zahlen +\def\bbbf{{\rm I\!F}} +\def\bbbh{{\rm I\!H}} +\def\bbbk{{\rm I\!K}} +\def\bbbp{{\rm I\!P}} +\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l} +{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}} +\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}} +\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm +Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}} +\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm +T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox +to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox +to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox +to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}} +\def\bbbs{{\mathchoice +{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox +to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox +to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox +to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox +to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}} +\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}} +{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}} +{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}} +{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}} + +\let\ts\, + +\setlength\leftmargini {17\p@} +\setlength\leftmargin {\leftmargini} +\setlength\leftmarginii {\leftmargini} +\setlength\leftmarginiii {\leftmargini} +\setlength\leftmarginiv {\leftmargini} +\setlength \labelsep {.5em} +\setlength \labelwidth{\leftmargini} +\addtolength\labelwidth{-\labelsep} + +\def\@listI{\leftmargin\leftmargini + \parsep 0\p@ \@plus1\p@ \@minus\p@ + \topsep 8\p@ \@plus2\p@ \@minus4\p@ + \itemsep0\p@} +\let\@listi\@listI +\@listi +\def\@listii {\leftmargin\leftmarginii + \labelwidth\leftmarginii + \advance\labelwidth-\labelsep + \topsep 0\p@ \@plus2\p@ \@minus\p@} +\def\@listiii{\leftmargin\leftmarginiii + \labelwidth\leftmarginiii + \advance\labelwidth-\labelsep + \topsep 0\p@ \@plus\p@\@minus\p@ + \parsep \z@ + \partopsep \p@ \@plus\z@ \@minus\p@} + +\renewcommand\labelitemi{\normalfont\bfseries --} +\renewcommand\labelitemii{$\m@th\bullet$} + +\setlength\arraycolsep{1.4\p@} +\setlength\tabcolsep{1.4\p@} + +\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}% + {{\contentsname}}} + \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}} + \def\lastand{\ifnum\value{auco}=2\relax + \unskip{} \andname\ + \else + \unskip \lastandname\ + \fi}% + \def\and{\stepcounter{@auth}\relax + \ifnum\value{@auth}=\value{auco}% + \lastand + \else + \unskip, + \fi}% + \@starttoc{toc}\if@restonecol\twocolumn\fi} + +\def\l@part#1#2{\addpenalty{\@secpenalty}% + \addvspace{2em plus\p@}% % space above part line + \begingroup + \parindent \z@ + \rightskip \z@ plus 5em + \hrule\vskip5pt + \large % same size as for a contribution heading + \bfseries\boldmath % set line in boldface + \leavevmode % TeX command to enter horizontal mode. + #1\par + \vskip5pt + \hrule + \vskip1pt + \nobreak % Never break after part entry + \endgroup} + +\def\@dotsep{2} + +\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else +{chapter.\thechapter}\fi} + +\def\addnumcontentsmark#1#2#3{% +\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline + {\thechapter}#3}{\thepage}\hyperhrefextend}} +\def\addcontentsmark#1#2#3{% +\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}} +\def\addcontentsmarkwop#1#2#3{% +\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}} + +\def\@adcmk[#1]{\ifcase #1 \or +\def\@gtempa{\addnumcontentsmark}% + \or \def\@gtempa{\addcontentsmark}% + \or \def\@gtempa{\addcontentsmarkwop}% + \fi\@gtempa{toc}{chapter}} +\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}} + +\def\l@chapter#1#2{\addpenalty{-\@highpenalty} + \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup + \parindent \z@ \rightskip \@tocrmarg + \advance\rightskip by 0pt plus 2cm + \parfillskip -\rightskip \pretolerance=10000 + \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip + {\large\bfseries\boldmath#1}\ifx0#2\hfil\null + \else + \nobreak + \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern + \@dotsep mu$}\hfill + \nobreak\hbox to\@pnumwidth{\hss #2}% + \fi\par + \penalty\@highpenalty \endgroup} + +\def\l@title#1#2{\addpenalty{-\@highpenalty} + \addvspace{8pt plus 1pt} + \@tempdima \z@ + \begingroup + \parindent \z@ \rightskip \@tocrmarg + \advance\rightskip by 0pt plus 2cm + \parfillskip -\rightskip \pretolerance=10000 + \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip + #1\nobreak + \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern + \@dotsep mu$}\hfill + \nobreak\hbox to\@pnumwidth{\hss #2}\par + \penalty\@highpenalty \endgroup} + +\def\l@author#1#2{\addpenalty{\@highpenalty} + \@tempdima=\z@ %15\p@ + \begingroup + \parindent \z@ \rightskip \@tocrmarg + \advance\rightskip by 0pt plus 2cm + \pretolerance=10000 + \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip + \textit{#1}\par + \penalty\@highpenalty \endgroup} + +%\setcounter{tocdepth}{0} +\newdimen\tocchpnum +\newdimen\tocsecnum +\newdimen\tocsectotal +\newdimen\tocsubsecnum +\newdimen\tocsubsectotal +\newdimen\tocsubsubsecnum +\newdimen\tocsubsubsectotal +\newdimen\tocparanum +\newdimen\tocparatotal +\newdimen\tocsubparanum +\tocchpnum=\z@ % no chapter numbers +\tocsecnum=15\p@ % section 88. plus 2.222pt +\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt +\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt +\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt +\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt +\def\calctocindent{% +\tocsectotal=\tocchpnum +\advance\tocsectotal by\tocsecnum +\tocsubsectotal=\tocsectotal +\advance\tocsubsectotal by\tocsubsecnum +\tocsubsubsectotal=\tocsubsectotal +\advance\tocsubsubsectotal by\tocsubsubsecnum +\tocparatotal=\tocsubsubsectotal +\advance\tocparatotal by\tocparanum} +\calctocindent + +\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}} +\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}} +\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}} +\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}} +\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}} + +\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn + \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}} + \@starttoc{lof}\if@restonecol\twocolumn\fi} +\def\l@figure{\@dottedtocline{1}{0em}{1.5em}} + +\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn + \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}} + \@starttoc{lot}\if@restonecol\twocolumn\fi} +\let\l@table\l@figure + +\renewcommand\listoffigures{% + \section*{\listfigurename + \@mkboth{\listfigurename}{\listfigurename}}% + \@starttoc{lof}% + } + +\renewcommand\listoftables{% + \section*{\listtablename + \@mkboth{\listtablename}{\listtablename}}% + \@starttoc{lot}% + } + +\ifx\oribibl\undefined +\ifx\citeauthoryear\undefined +\renewenvironment{thebibliography}[1] + {\section*{\refname} + \def\@biblabel##1{##1.} + \small + \list{\@biblabel{\@arabic\c@enumiv}}% + {\settowidth\labelwidth{\@biblabel{#1}}% + \leftmargin\labelwidth + \advance\leftmargin\labelsep + \if@openbib + \advance\leftmargin\bibindent + \itemindent -\bibindent + \listparindent \itemindent + \parsep \z@ + \fi + \usecounter{enumiv}% + \let\p@enumiv\@empty + \renewcommand\theenumiv{\@arabic\c@enumiv}}% + \if@openbib + \renewcommand\newblock{\par}% + \else + \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}% + \fi + \sloppy\clubpenalty4000\widowpenalty4000% + \sfcode`\.=\@m} + {\def\@noitemerr + {\@latex@warning{Empty `thebibliography' environment}}% + \endlist} +\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw + {\let\protect\noexpand\immediate + \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces} +\newcount\@tempcntc +\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi + \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do + {\@ifundefined + {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries + ?}\@warning + {Citation `\@citeb' on page \thepage \space undefined}}% + {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}% + \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne + \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}% + \else + \advance\@tempcntb\@ne + \ifnum\@tempcntb=\@tempcntc + \else\advance\@tempcntb\m@ne\@citeo + \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}} +\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else + \@citea\def\@citea{,\,\hskip\z@skip}% + \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else + {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else + \def\@citea{--}\fi + \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi} +\else +\renewenvironment{thebibliography}[1] + {\section*{\refname} + \small + \list{}% + {\settowidth\labelwidth{}% + \leftmargin\parindent + \itemindent=-\parindent + \labelsep=\z@ + \if@openbib + \advance\leftmargin\bibindent + \itemindent -\bibindent + \listparindent \itemindent + \parsep \z@ + \fi + \usecounter{enumiv}% + \let\p@enumiv\@empty + \renewcommand\theenumiv{}}% + \if@openbib + \renewcommand\newblock{\par}% + \else + \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}% + \fi + \sloppy\clubpenalty4000\widowpenalty4000% + \sfcode`\.=\@m} + {\def\@noitemerr + {\@latex@warning{Empty `thebibliography' environment}}% + \endlist} + \def\@cite#1{#1}% + \def\@lbibitem[#1]#2{\item[]\if@filesw + {\def\protect##1{\string ##1\space}\immediate + \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces} + \fi +\else +\@cons\@openbib@code{\noexpand\small} +\fi + +\def\idxquad{\hskip 10\p@}% space that divides entry from number + +\def\@idxitem{\par\hangindent 10\p@} + +\def\subitem{\par\setbox0=\hbox{--\enspace}% second order + \noindent\hangindent\wd0\box0}% index entry + +\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third + \noindent\hangindent\wd0\box0}% order index entry + +\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax} + +\renewenvironment{theindex} + {\@mkboth{\indexname}{\indexname}% + \thispagestyle{empty}\parindent\z@ + \parskip\z@ \@plus .3\p@\relax + \let\item\par + \def\,{\relax\ifmmode\mskip\thinmuskip + \else\hskip0.2em\ignorespaces\fi}% + \normalfont\small + \begin{multicols}{2}[\@makeschapterhead{\indexname}]% + } + {\end{multicols}} + +\renewcommand\footnoterule{% + \kern-3\p@ + \hrule\@width 2truecm + \kern2.6\p@} + \newdimen\fnindent + \fnindent1em +\long\def\@makefntext#1{% + \parindent \fnindent% + \leftskip \fnindent% + \noindent + \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1} + +\long\def\@makecaption#1#2{% + \vskip\abovecaptionskip + \sbox\@tempboxa{{\bfseries #1.} #2}% + \ifdim \wd\@tempboxa >\hsize + {\bfseries #1.} #2\par + \else + \global \@minipagefalse + \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}% + \fi + \vskip\belowcaptionskip} + +\def\fps@figure{htbp} +\def\fnum@figure{\figurename\thinspace\thefigure} +\def \@floatboxreset {% + \reset@font + \small + \@setnobreak + \@setminipage +} +\def\fps@table{htbp} +\def\fnum@table{\tablename~\thetable} +\renewenvironment{table} + {\setlength\abovecaptionskip{0\p@}% + \setlength\belowcaptionskip{10\p@}% + \@float{table}} + {\end@float} +\renewenvironment{table*} + {\setlength\abovecaptionskip{0\p@}% + \setlength\belowcaptionskip{10\p@}% + \@dblfloat{table}} + {\end@dblfloat} + +\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname + ext@#1\endcsname}{#1}{\protect\numberline{\csname + the#1\endcsname}{\ignorespaces #2}}\begingroup + \@parboxrestore + \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par + \endgroup} + +% LaTeX does not provide a command to enter the authors institute +% addresses. The \institute command is defined here. + +\newcounter{@inst} +\newcounter{@auth} +\newcounter{auco} +\newdimen\instindent +\newbox\authrun +\newtoks\authorrunning +\newtoks\tocauthor +\newbox\titrun +\newtoks\titlerunning +\newtoks\toctitle + +\def\clearheadinfo{\gdef\@author{No Author Given}% + \gdef\@title{No Title Given}% + \gdef\@subtitle{}% + \gdef\@institute{No Institute Given}% + \gdef\@thanks{}% + \global\titlerunning={}\global\authorrunning={}% + \global\toctitle={}\global\tocauthor={}} + +\def\institute#1{\gdef\@institute{#1}} + +\def\institutename{\par + \begingroup + \parskip=\z@ + \parindent=\z@ + \setcounter{@inst}{1}% + \def\and{\par\stepcounter{@inst}% + \noindent$^{\the@inst}$\enspace\ignorespaces}% + \setbox0=\vbox{\def\thanks##1{}\@institute}% + \ifnum\c@@inst=1\relax + \gdef\fnnstart{0}% + \else + \xdef\fnnstart{\c@@inst}% + \setcounter{@inst}{1}% + \noindent$^{\the@inst}$\enspace + \fi + \ignorespaces + \@institute\par + \endgroup} + +\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or + {\star\star\star}\or \dagger\or \ddagger\or + \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger + \or \ddagger\ddagger \else\@ctrerr\fi}} + +\def\inst#1{\unskip$^{#1}$} +\def\fnmsep{\unskip$^,$} +\def\email#1{{\tt#1}} +\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}% +\@ifpackageloaded{babel}{% +\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}% +\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}% +\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}% +}{\switcht@@therlang}% +} +\def\homedir{\~{ }} + +\def\subtitle#1{\gdef\@subtitle{#1}} +\clearheadinfo + +\renewcommand\maketitle{\newpage + \refstepcounter{chapter}% + \stepcounter{section}% + \setcounter{section}{0}% + \setcounter{subsection}{0}% + \setcounter{figure}{0} + \setcounter{table}{0} + \setcounter{equation}{0} + \setcounter{footnote}{0}% + \begingroup + \parindent=\z@ + \renewcommand\thefootnote{\@fnsymbol\c@footnote}% + \if@twocolumn + \ifnum \col@number=\@ne + \@maketitle + \else + \twocolumn[\@maketitle]% + \fi + \else + \newpage + \global\@topnum\z@ % Prevents figures from going at top of page. + \@maketitle + \fi + \thispagestyle{empty}\@thanks +% + \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}% + \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}% + \instindent=\hsize + \advance\instindent by-\headlineindent +% \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else +% \addcontentsline{toc}{title}{\the\toctitle}\fi + \if@runhead + \if!\the\titlerunning!\else + \edef\@title{\the\titlerunning}% + \fi + \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}% + \ifdim\wd\titrun>\instindent + \typeout{Title too long for running head. Please supply}% + \typeout{a shorter form with \string\titlerunning\space prior to + \string\maketitle}% + \global\setbox\titrun=\hbox{\small\rm + Title Suppressed Due to Excessive Length}% + \fi + \xdef\@title{\copy\titrun}% + \fi +% + \if!\the\tocauthor!\relax + {\def\and{\noexpand\protect\noexpand\and}% + \protected@xdef\toc@uthor{\@author}}% + \else + \def\\{\noexpand\protect\noexpand\newline}% + \protected@xdef\scratch{\the\tocauthor}% + \protected@xdef\toc@uthor{\scratch}% + \fi +% \addcontentsline{toc}{author}{\toc@uthor}% + \if@runhead + \if!\the\authorrunning! + \value{@inst}=\value{@auth}% + \setcounter{@auth}{1}% + \else + \edef\@author{\the\authorrunning}% + \fi + \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}% + \ifdim\wd\authrun>\instindent + \typeout{Names of authors too long for running head. Please supply}% + \typeout{a shorter form with \string\authorrunning\space prior to + \string\maketitle}% + \global\setbox\authrun=\hbox{\small\rm + Authors Suppressed Due to Excessive Length}% + \fi + \xdef\@author{\copy\authrun}% + \markboth{\@author}{\@title}% + \fi + \endgroup + \setcounter{footnote}{\fnnstart}% + \clearheadinfo} +% +\def\@maketitle{\newpage + \markboth{}{}% + \def\lastand{\ifnum\value{@inst}=2\relax + \unskip{} \andname\ + \else + \unskip \lastandname\ + \fi}% + \def\and{\stepcounter{@auth}\relax + \ifnum\value{@auth}=\value{@inst}% + \lastand + \else + \unskip, + \fi}% + \begin{center}% + \let\newline\\ + {\Large \bfseries\boldmath + \pretolerance=10000 + \@title \par}\vskip .8cm +\if!\@subtitle!\else {\large \bfseries\boldmath + \vskip -.65cm + \pretolerance=10000 + \@subtitle \par}\vskip .8cm\fi + \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}% + \def\thanks##1{}\@author}% + \global\value{@inst}=\value{@auth}% + \global\value{auco}=\value{@auth}% + \setcounter{@auth}{1}% +{\lineskip .5em +\noindent\ignorespaces +\@author\vskip.35cm} + {\small\institutename} + \end{center}% + } + +% definition of the "\spnewtheorem" command. +% +% Usage: +% +% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font} +% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font} +% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font} +% +% New is "cap_font" and "body_font". It stands for +% fontdefinition of the caption and the text itself. +% +% "\spnewtheorem*" gives a theorem without number. +% +% A defined spnewthoerem environment is used as described +% by Lamport. +% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +\def\@thmcountersep{} +\def\@thmcounterend{.} + +\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}} + +% definition of \spnewtheorem with number + +\def\@spnthm#1#2{% + \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}} +\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}} + +\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname + {\@definecounter{#1}\@addtoreset{#1}{#3}% + \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand + \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}% + \expandafter\xdef\csname #1name\endcsname{#2}% + \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}% + \global\@namedef{end#1}{\@endtheorem}}} + +\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname + {\@definecounter{#1}% + \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}% + \expandafter\xdef\csname #1name\endcsname{#2}% + \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}% + \global\@namedef{end#1}{\@endtheorem}}} + +\def\@spothm#1[#2]#3#4#5{% + \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}% + {\expandafter\@ifdefinable\csname #1\endcsname + {\global\@namedef{the#1}{\@nameuse{the#2}}% + \expandafter\xdef\csname #1name\endcsname{#3}% + \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}% + \global\@namedef{end#1}{\@endtheorem}}}} + +\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@ +\refstepcounter{#1}% +\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}} + +\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}% + \ignorespaces} + +\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname + the#1\endcsname}{#5}{#3}{#4}\ignorespaces} + +\def\@spbegintheorem#1#2#3#4{\trivlist + \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4} + +\def\@spopargbegintheorem#1#2#3#4#5{\trivlist + \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5} + +% definition of \spnewtheorem* without number + +\def\@sthm#1#2{\@Ynthm{#1}{#2}} + +\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname + {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}% + \expandafter\xdef\csname #1name\endcsname{#2}% + \global\@namedef{end#1}{\@endtheorem}}} + +\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@ +\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}} + +\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces} + +\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1} + {#4}{#2}{#3}\ignorespaces} + +\def\@Begintheorem#1#2#3{#3\trivlist + \item[\hskip\labelsep{#2#1\@thmcounterend}]} + +\def\@Opargbegintheorem#1#2#3#4{#4\trivlist + \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }} + +\if@envcntsect + \def\@thmcountersep{.} + \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape} +\else + \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape} + \if@envcntreset + \@addtoreset{theorem}{section} + \else + \@addtoreset{theorem}{chapter} + \fi +\fi + +%definition of divers theorem environments +\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily} +\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily} +\if@envcntsame % alle Umgebungen wie Theorem. + \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}} +\else % alle Umgebungen mit eigenem Zaehler + \if@envcntsect % mit section numeriert + \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}} + \else % nicht mit section numeriert + \if@envcntreset + \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4} + \@addtoreset{#1}{section}} + \else + \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4} + \@addtoreset{#1}{chapter}}% + \fi + \fi +\fi +\spn@wtheorem{case}{Case}{\itshape}{\rmfamily} +\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily} +\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape} +\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape} +\spn@wtheorem{example}{Example}{\itshape}{\rmfamily} +\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily} +\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape} +\spn@wtheorem{note}{Note}{\itshape}{\rmfamily} +\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily} +\spn@wtheorem{property}{Property}{\itshape}{\rmfamily} +\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape} +\spn@wtheorem{question}{Question}{\itshape}{\rmfamily} +\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily} +\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily} + +\def\@takefromreset#1#2{% + \def\@tempa{#1}% + \let\@tempd\@elt + \def\@elt##1{% + \def\@tempb{##1}% + \ifx\@tempa\@tempb\else + \@addtoreset{##1}{#2}% + \fi}% + \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname + \expandafter\def\csname cl@#2\endcsname{}% + \@tempc + \let\@elt\@tempd} + +\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist + \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5} + \def\@Opargbegintheorem##1##2##3##4{##4\trivlist + \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }} + } + +\renewenvironment{abstract}{% + \list{}{\advance\topsep by0.35cm\relax\small + \leftmargin=1cm + \labelwidth=\z@ + \listparindent=\z@ + \itemindent\listparindent + \rightmargin\leftmargin}\item[\hskip\labelsep + \bfseries\abstractname]} + {\endlist} + +\newdimen\headlineindent % dimension for space between +\headlineindent=1.166cm % number and text of headings. + +\def\ps@headings{\let\@mkboth\@gobbletwo + \let\@oddfoot\@empty\let\@evenfoot\@empty + \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}% + \leftmark\hfil} + \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}% + \llap{\thepage}} + \def\chaptermark##1{}% + \def\sectionmark##1{}% + \def\subsectionmark##1{}} + +\def\ps@titlepage{\let\@mkboth\@gobbletwo + \let\@oddfoot\@empty\let\@evenfoot\@empty + \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}% + \hfil} + \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}% + \llap{\thepage}} + \def\chaptermark##1{}% + \def\sectionmark##1{}% + \def\subsectionmark##1{}} + +\if@runhead\ps@headings\else +\ps@empty\fi + +\setlength\arraycolsep{1.4\p@} +\setlength\tabcolsep{1.4\p@} + +\endinput +%end of file llncs.cls