# HG changeset patch # User wenzelm # Date 1450802506 -3600 # Node ID ad710de5c57664eacb07bc84aeff9badd6e92f23 # Parent fe2b7f4276be88297ee099be9ca013401022d35a more standard nesting of sub-language: Parse.text allows atomic entities without quotes; diff -r fe2b7f4276be -r ad710de5c576 src/Doc/Eisbach/Manual.thy --- a/src/Doc/Eisbach/Manual.thy Tue Dec 22 17:14:35 2015 +0100 +++ b/src/Doc/Eisbach/Manual.thy Tue Dec 22 17:41:46 2015 +0100 @@ -284,7 +284,7 @@ further defined in @{cite "isabelle-isar-ref"}. @{rail \ - @@{method match} kind @'in' (pattern '\' cartouche + '\') + @@{method match} kind @'in' (pattern '\' @{syntax text} + '\') ; kind: (@'conclusion' | @'premises' ('(' 'local' ')')? | diff -r fe2b7f4276be -r ad710de5c576 src/HOL/Eisbach/Examples.thy --- a/src/HOL/Eisbach/Examples.thy Tue Dec 22 17:14:35 2015 +0100 +++ b/src/HOL/Eisbach/Examples.thy Tue Dec 22 17:41:46 2015 +0100 @@ -154,7 +154,7 @@ (assumption | rule intros | erule elims | subst subst | subst (asm) subst | - (erule notE; solves \prop_solver\))+ + (erule notE; solves prop_solver))+ lemmas [intros] = conjI @@ -199,7 +199,7 @@ done method fol_solver = - ((guess_ex | guess_all | prop_solver) ; solves \fol_solver\) + ((guess_ex | guess_all | prop_solver); solves fol_solver) declare allI [intros] @@ -241,7 +241,7 @@ assumes imps: "A \ B \ C" "D \ C" "E \ D \ A" shows "(A \ B \ C) \ (D \ C)" by (match imps[uncurry] in H[curry]:"_ \ C" (cut, multi) \ - \match H in "E \ _" \ \fail\ + \match H in "E \ _" \ fail \ _ \ \simp add: H\\) end diff -r fe2b7f4276be -r ad710de5c576 src/HOL/Eisbach/Tests.thy --- a/src/HOL/Eisbach/Tests.thy Tue Dec 22 17:14:35 2015 +0100 +++ b/src/HOL/Eisbach/Tests.thy Tue Dec 22 17:41:46 2015 +0100 @@ -21,7 +21,7 @@ apply foo done -method abs_used for P = (match (P) in "\a. ?Q" \ \fail\ \ _ \ \-\) +method abs_used for P = (match (P) in "\a. ?Q" \ fail \ _ \ -) subsection \Match Tests\ @@ -74,10 +74,10 @@ fix C have "(\x :: 'a. A x) \ (\z. B z) \ A y \ B x \ B x \ A y" apply (intro conjI) - apply (match premises in Y: "\z :: 'a. A z" and Y'[intro]:"\z :: 'b. B z" \ \fastforce\) - apply (match premises in Y: "\z :: 'a. A z" \ \match premises in Y'[intro]:"\z :: 'b. B z" \ \fastforce\\) + apply (match premises in Y: "\z :: 'a. A z" and Y'[intro]:"\z :: 'b. B z" \ fastforce) + apply (match premises in Y: "\z :: 'a. A z" \ \match premises in Y'[intro]:"\z :: 'b. B z" \ fastforce\) apply (match premises in Y[thin]: "\z :: 'a. A z" \ \(match premises in Y':"\z :: 'a. A z" \ \print_fact Y,fail\ \ _ \ \print_fact Y\)\) - (*apply (match premises in Y: "\z :: 'b. B z" \ \(match premises in Y'[thin]:"\z :: 'b. B z" \ \(match premises in Y':"\z :: 'a. A z" \ \fail\ \ Y': _ \ \-\)\)\)*) + (*apply (match premises in Y: "\z :: 'b. B z" \ \(match premises in Y'[thin]:"\z :: 'b. B z" \ \(match premises in Y':"\z :: 'a. A z" \ fail \ Y': _ \ -)\)\)*) apply assumption done @@ -97,14 +97,14 @@ fix A B C x have "(\x :: 'a. A x) \ (\y :: 'a. B y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y[thin]: "\z :: 'a. ?A z" (multi) \ \intro conjI, (rule Y)+\) - apply (match premises in Y[thin]: "\z :: 'a. ?A z" (multi) \ \fail\ \ "C y" \ \-\) (* multi-match must bind something *) + apply (match premises in Y[thin]: "\z :: 'a. ?A z" (multi) \ fail \ "C y" \ -) (* multi-match must bind something *) apply (match premises in Y: "C y" \ \rule Y\) done fix A B C x have "(\x :: 'a. A x) \ (\y :: 'a. B y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y[thin]: "\z. ?A z" (multi) \ \intro conjI, (rule Y)+\) - apply (match premises in Y[thin]: "\z. ?A z" (multi) \ \fail\ \ "C y" \ \-\) (* multi-match must bind something *) + apply (match premises in Y[thin]: "\z. ?A z" (multi) \ fail \ "C y" \ -) (* multi-match must bind something *) apply (match premises in Y: "C y" \ \rule Y\) done @@ -117,9 +117,9 @@ fix A B C x have "(\y :: 'a. B y \ C y) \ (\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ - \match (P) in B \ \fail\ - \ "\a. B" \ \fail\ - \ _ \ \-\, + \match (P) in B \ fail + \ "\a. B" \ fail + \ _ \ -, intro conjI, (rule Y[THEN conjunct1])\) apply (rule dup) apply (match premises in Y':"\x :: 'a. ?U x \ Q x" and Y: "\x :: 'a. Q x \ ?U x" (multi) for Q \ \insert Y[THEN conjunct1]\) @@ -139,7 +139,7 @@ apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where a=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where a=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where z=x,THEN conjunct1]\) - apply (match asms in Y: "\z a. A (z :: 'a) (a :: 'a) \ R" for R \ \fail\ \ _ \ \-\) + apply (match asms in Y: "\z a. A (z :: 'a) (a :: 'a) \ R" for R \ fail \ _ \ -) apply (rule asms[THEN conjunct1]) done @@ -188,7 +188,7 @@ assume X: "A \ B" "A \ C" C have "A \ B \ C" by (match X in H: "A \ ?H" (multi, cut) \ - \match H in "A \ C" and "A \ B" \ \fail\\ + \match H in "A \ C" and "A \ B" \ fail\ | simp add: X) @@ -199,10 +199,10 @@ have "(\x. A x \ B) \ B \ C \ C" apply (match premises in H:"\x. A x \ B" \ \match premises in H'[thin]: "\x. A x \ B" \ - \match premises in H'':"\x. A x \ B" \ \fail\ - \ _ \ \-\\ - ,match premises in H'':"\x. A x \ B" \ \fail\ \ _ \ \-\\) - apply (match premises in H:"\x. A x \ B" \ \fail\ + \match premises in H'':"\x. A x \ B" \ fail + \ _ \ -\ + ,match premises in H'':"\x. A x \ B" \ fail \ _ \ -\) + apply (match premises in H:"\x. A x \ B" \ fail \ H':_ \ \rule H'[THEN conjunct2]\) done @@ -214,7 +214,7 @@ assume asms: "C \ D" have "B \ C \ C" by (match premises in _ \ \insert asms, - match premises (local) in "B \ C" \ \fail\ + match premises (local) in "B \ C" \ fail \ H:"C \ D" \ \rule H[THEN conjunct1]\\) end @@ -233,13 +233,13 @@ lemma assumes A: "P \ Q" shows "P" - by ((match A in "P \ Q" \ \fail\ \ "?H" \ \-\) | simp add: A) + by ((match A in "P \ Q" \ fail \ "?H" \ -) | simp add: A) lemma assumes A: "D \ C" "A \ B" "A \ B" shows "A" apply ((match A in U: "P \ Q" (cut) and "P' \ Q'" for P Q P' Q' \ - \simp add: U\ \ "?H" \ \-\) | -) + \simp add: U\ \ "?H" \ -) | -) apply (simp add: A) done @@ -283,8 +283,8 @@ (match facts in H: "A" and H': "B" \ \recursion_example' "A" "B" facts: H TrueI\ \ "A" and "True" \ \recursion_example' "A" "B" facts: TrueI\ - \ "True" \ \-\ - \ "PROP ?P" \ \fail\) + \ "True" \ - + \ "PROP ?P" \ fail) lemma assumes asms: "A" "B" @@ -460,9 +460,9 @@ subsection \Proper context for method parameters\ -method add_simp methods m uses f = (match f in H[simp]:_ \ \m\) +method add_simp methods m uses f = (match f in H[simp]:_ \ m) -method add_my_thms methods m uses f = (match f in H[my_thms_named]:_ \ \m\) +method add_my_thms methods m uses f = (match f in H[my_thms_named]:_ \ m) method rule_my_thms = (rule my_thms_named) method rule_my_thms' declares my_thms_named = (rule my_thms_named) @@ -472,9 +472,9 @@ shows "(A \ B) \ A \ A \ A" apply (intro conjI) - apply (add_simp \add_simp \simp\ f: B\ f: A) - apply (add_my_thms \rule_my_thms\ f:A) - apply (add_my_thms \rule_my_thms'\ f:A) + apply (add_simp \add_simp simp f: B\ f: A) + apply (add_my_thms rule_my_thms f:A) + apply (add_my_thms rule_my_thms' f:A) apply (add_my_thms \rule my_thms_named\ f:A) done @@ -483,7 +483,7 @@ method all_args for A B methods m1 m2 uses f1 f2 declares my_thms_named = fail lemma True - by (all_args True False \-\ \fail\ f1: TrueI f2: TrueI my_thms_named: TrueI | rule TrueI) + by (all_args True False - fail f1: TrueI f2: TrueI my_thms_named: TrueI | rule TrueI) subsection \Method name internalization test\ diff -r fe2b7f4276be -r ad710de5c576 src/HOL/Eisbach/match_method.ML --- a/src/HOL/Eisbach/match_method.ML Tue Dec 22 17:14:35 2015 +0100 +++ b/src/HOL/Eisbach/match_method.ML Tue Dec 22 17:41:46 2015 +0100 @@ -94,9 +94,9 @@ else let val b = #1 (the opt_dyn) in error ("Cannot bind fact name in term match" ^ Position.here (pos_of b)) end)) -- - Scan.lift (for_fixes -- (@{keyword "\"} |-- Parse.token Parse.cartouche)) - >> (fn ((ctxt, ts), (fixes, cartouche)) => - (case Token.get_value cartouche of + Scan.lift (for_fixes -- (@{keyword "\"} |-- Parse.token Parse.text)) + >> (fn ((ctxt, ts), (fixes, body)) => + (case Token.get_value body of SOME (Token.Source src) => let val text = Method_Closure.read ctxt src; @@ -198,7 +198,7 @@ |> (fn ctxt => fold2 upd_ctxt binds pats ([], ctxt) |> apfst rev) ||> Proof_Context.restore_mode ctxt; - val (text, src) = Method_Closure.read_closure_input ctxt6 (Token.input_of cartouche); + val (text, src) = Method_Closure.read_closure_input ctxt6 (Token.input_of body); val morphism = Variable.export_morphism ctxt6 @@ -239,7 +239,7 @@ val binds'' = (binds' ~~ match_args) ~~ pats'; val src' = map (Token.transform morphism) src; - val _ = Token.assign (SOME (Token.Source src')) cartouche; + val _ = Token.assign (SOME (Token.Source src')) body; in (binds'', real_fixes', text) end)); diff -r fe2b7f4276be -r ad710de5c576 src/HOL/Eisbach/method_closure.ML --- a/src/HOL/Eisbach/method_closure.ML Tue Dec 22 17:14:35 2015 +0100 +++ b/src/HOL/Eisbach/method_closure.ML Tue Dec 22 17:41:46 2015 +0100 @@ -90,7 +90,7 @@ val method_text = - Args.context -- Scan.lift (Parse.token Parse.cartouche) >> (fn (ctxt, tok) => + Args.context -- Scan.lift (Parse.token Parse.text) >> (fn (ctxt, tok) => (case Token.get_value tok of SOME (Token.Source src) => read ctxt src | _ =>