# HG changeset patch # User wenzelm # Date 1451335652 -3600 # Node ID e96292f32c3cfae42a5941dabdb21d2ae4d25a6c # Parent 1d43f86f48be3f67a85df6354fdb3a4e5fd5f397 former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII"; diff -r 1d43f86f48be -r e96292f32c3c NEWS --- a/NEWS Mon Dec 28 19:23:15 2015 +0100 +++ b/NEWS Mon Dec 28 21:47:32 2015 +0100 @@ -375,6 +375,11 @@ *** HOL *** +* Former "xsymbols" syntax with Isabelle symbols is used by default, +without any special print mode. Important ASCII replacement syntax +remains available under print mode "ASCII", but less important syntax +has been removed (see below). + * Combinator to represent case distinction on products is named "case_prod", uniformly, discontinuing any input aliasses. Very popular theorem aliasses have been retained. diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Complete_Lattices.thy Mon Dec 28 21:47:32 2015 +0100 @@ -85,27 +85,27 @@ with the plain constant names. \ -syntax +syntax (ASCII) "_INF1" :: "pttrns \ 'b \ 'b" ("(3INF _./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3INF _:_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3SUP _./ _)" [0, 10] 10) "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3SUP _:_./ _)" [0, 0, 10] 10) -syntax (xsymbols) +syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) translations - "INF x y. B" == "INF x. INF y. B" - "INF x. B" == "CONST INFIMUM CONST UNIV (%x. B)" - "INF x. B" == "INF x:CONST UNIV. B" - "INF x:A. B" == "CONST INFIMUM A (%x. B)" - "SUP x y. B" == "SUP x. SUP y. B" - "SUP x. B" == "CONST SUPREMUM CONST UNIV (%x. B)" - "SUP x. B" == "SUP x:CONST UNIV. B" - "SUP x:A. B" == "CONST SUPREMUM A (%x. B)" + "\x y. B" \ "\x. \y. B" + "\x. B" \ "CONST INFIMUM CONST UNIV (\x. B)" + "\x. B" \ "\x \ CONST UNIV. B" + "\x\A. B" \ "CONST INFIMUM A (\x. B)" + "\x y. B" \ "\x. \y. B" + "\x. B" \ "CONST SUPREMUM CONST UNIV (\x. B)" + "\x. B" \ "\x \ CONST UNIV. B" + "\x\A. B" \ "CONST SUPREMUM A (\x. B)" print_translation \ [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFIMUM} @{syntax_const "_INF"}, @@ -970,31 +970,31 @@ subsubsection \Intersections of families\ -abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where - "INTER \ INFIMUM" +abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" + where "INTER \ INFIMUM" text \ Note: must use name @{const INTER} here instead of \INT\ to allow the following syntax coexist with the plain constant name. \ -syntax - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) - -syntax (xsymbols) - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) +syntax (ASCII) + "_INTER1" :: "pttrns \ 'b set \ 'b set" ("(3INT _./ _)" [0, 10] 10) + "_INTER" :: "pttrn \ 'a set \ 'b set \ 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) syntax (latex output) - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) + "_INTER1" :: "pttrns \ 'b set \ 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "_INTER" :: "pttrn \ 'a set \ 'b set \ 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) + +syntax + "_INTER1" :: "pttrns \ 'b set \ 'b set" ("(3\_./ _)" [0, 10] 10) + "_INTER" :: "pttrn \ 'a set \ 'b set \ 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) translations - "INT x y. B" == "INT x. INT y. B" - "INT x. B" == "CONST INTER CONST UNIV (%x. B)" - "INT x. B" == "INT x:CONST UNIV. B" - "INT x:A. B" == "CONST INTER A (%x. B)" + "\x y. B" \ "\x. \y. B" + "\x. B" \ "CONST INTER CONST UNIV (\x. B)" + "\x. B" \ "\x \ CONST UNIV. B" + "\x\A. B" \ "CONST INTER A (\x. B)" print_translation \ [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}] @@ -1143,39 +1143,36 @@ subsubsection \Unions of families\ -abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" where - "UNION \ SUPREMUM" +abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" + where "UNION \ SUPREMUM" text \ Note: must use name @{const UNION} here instead of \UN\ to allow the following syntax coexist with the plain constant name. \ -syntax +syntax (ASCII) "_UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 0, 10] 10) -syntax (xsymbols) - "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) - syntax (latex output) "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) +syntax + "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) + translations - "UN x y. B" == "UN x. UN y. B" - "UN x. B" == "CONST UNION CONST UNIV (%x. B)" - "UN x. B" == "UN x:CONST UNIV. B" - "UN x:A. B" == "CONST UNION A (%x. B)" + "\x y. B" \ "\x. \y. B" + "\x. B" \ "CONST UNION CONST UNIV (\x. B)" + "\x. B" \ "\x \ CONST UNIV. B" + "\x\A. B" \ "CONST UNION A (\x. B)" text \ - Note the difference between ordinary xsymbol syntax of indexed + Note the difference between ordinary syntax of indexed unions and intersections (e.g.\ \\a\<^sub>1\A\<^sub>1. B\) - and their \LaTeX\ rendition: @{term"\a\<^sub>1\A\<^sub>1. B"}. The - former does not make the index expression a subscript of the - union/intersection symbol because this leads to problems with nested - subscripts in Proof General. + and their \LaTeX\ rendition: @{term"\a\<^sub>1\A\<^sub>1. B"}. \ print_translation \ diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Filter.thy --- a/src/HOL/Filter.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Filter.thy Mon Dec 28 21:47:32 2015 +0100 @@ -693,19 +693,20 @@ lemma eventually_sequentially_seg: "eventually (\n. P (n + k)) sequentially \ eventually P sequentially" using eventually_sequentially_Suc[of "\n. P (n + k)" for k] by (induction k) auto -subsection \ The cofinite filter \ + +subsection \The cofinite filter\ definition "cofinite = Abs_filter (\P. finite {x. \ P x})" -abbreviation Inf_many :: "('a \ bool) \ bool" (binder "INFM " 10) where - "Inf_many P \ frequently P cofinite" +abbreviation Inf_many :: "('a \ bool) \ bool" (binder "\\<^sub>\" 10) + where "Inf_many P \ frequently P cofinite" -abbreviation Alm_all :: "('a \ bool) \ bool" (binder "MOST " 10) where - "Alm_all P \ eventually P cofinite" +abbreviation Alm_all :: "('a \ bool) \ bool" (binder "\\<^sub>\" 10) + where "Alm_all P \ eventually P cofinite" -notation (xsymbols) - Inf_many (binder "\\<^sub>\" 10) and - Alm_all (binder "\\<^sub>\" 10) +notation (ASCII) + Inf_many (binder "INFM " 10) and + Alm_all (binder "MOST " 10) lemma eventually_cofinite: "eventually P cofinite \ finite {x. \ P x}" unfolding cofinite_def diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Fun.thy --- a/src/HOL/Fun.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Fun.thy Mon Dec 28 21:47:32 2015 +0100 @@ -42,11 +42,11 @@ subsection \The Composition Operator \f \ g\\ -definition comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "o" 55) where - "f o g = (\x. f (g x))" +definition comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "\" 55) + where "f \ g = (\x. f (g x))" -notation (xsymbols) - comp (infixl "\" 55) +notation (ASCII) + comp (infixl "o" 55) lemma comp_apply [simp]: "(f o g) x = f (g x)" by (simp add: comp_def) diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Groups_Big.thy Mon Dec 28 21:47:32 2015 +0100 @@ -471,35 +471,28 @@ defines setsum = setsum.F .. -abbreviation - Setsum ("\_" [1000] 999) where - "\A \ setsum (%x. x) A" +abbreviation Setsum ("\_" [1000] 999) + where "\A \ setsum (\x. x) A" end -text\Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is -written \\x\A. e\.\ +text \Now: lot's of fancy syntax. First, @{term "setsum (\x. e) A"} is written \\x\A. e\.\ +syntax (ASCII) + "_setsum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(3SUM _:_./ _)" [0, 51, 10] 10) syntax - "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_./ _)" [0, 51, 10] 10) -syntax (xsymbols) - "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(2\_\_./ _)" [0, 51, 10] 10) - + "_setsum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(2\_\_./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ - "SUM i:A. b" == "CONST setsum (%i. b) A" - "\i\A. b" == "CONST setsum (%i. b) A" + "\i\A. b" \ "CONST setsum (\i. b) A" -text\Instead of @{term"\x\{x. P}. e"} we introduce the shorter - \\x|P. e\.\ +text \Instead of @{term"\x\{x. P}. e"} we introduce the shorter \\x|P. e\.\ +syntax (ASCII) + "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0, 0, 10] 10) syntax - "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10) -syntax (xsymbols) - "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0,0,10] 10) - + "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations - "SUM x|P. t" => "CONST setsum (%x. t) {x. P}" - "\x|P. t" => "CONST setsum (%x. t) {x. P}" + "\x|P. t" => "CONST setsum (\x. t) {x. P}" print_translation \ let @@ -1059,26 +1052,21 @@ end -syntax +syntax (ASCII) "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(4PROD _:_./ _)" [0, 51, 10] 10) -syntax (xsymbols) +syntax "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(2\_\_./ _)" [0, 51, 10] 10) - translations \ \Beware of argument permutation!\ - "PROD i:A. b" == "CONST setprod (%i. b) A" - "\i\A. b" == "CONST setprod (%i. b) A" + "\i\A. b" == "CONST setprod (\i. b) A" -text\Instead of @{term"\x\{x. P}. e"} we introduce the shorter - \\x|P. e\.\ +text \Instead of @{term"\x\{x. P}. e"} we introduce the shorter \\x|P. e\.\ +syntax (ASCII) + "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(4PROD _ |/ _./ _)" [0, 0, 10] 10) syntax - "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(4PROD _ |/ _./ _)" [0,0,10] 10) -syntax (xsymbols) - "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0,0,10] 10) - + "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations - "PROD x|P. t" => "CONST setprod (%x. t) {x. P}" - "\x|P. t" => "CONST setprod (%x. t) {x. P}" + "\x|P. t" => "CONST setprod (\x. t) {x. P}" context comm_monoid_mult begin diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Groups_List.thy --- a/src/HOL/Groups_List.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Groups_List.thy Mon Dec 28 21:47:32 2015 +0100 @@ -93,15 +93,12 @@ end text \Some syntactic sugar for summing a function over a list:\ - +syntax (ASCII) + "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10) syntax - "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10) -syntax (xsymbols) "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\_\_. _)" [0, 51, 10] 10) - translations \ \Beware of argument permutation!\ - "SUM x<-xs. b" == "CONST listsum (CONST map (%x. b) xs)" - "\x\xs. b" == "CONST listsum (CONST map (%x. b) xs)" + "\x\xs. b" == "CONST listsum (CONST map (\x. b) xs)" text \TODO duplicates\ lemmas listsum_simps = listsum.Nil listsum.Cons @@ -352,13 +349,11 @@ text \Some syntactic sugar:\ -syntax +syntax (ASCII) "_listprod" :: "pttrn => 'a list => 'b => 'b" ("(3PROD _<-_. _)" [0, 51, 10] 10) -syntax (xsymbols) +syntax "_listprod" :: "pttrn => 'a list => 'b => 'b" ("(3\_\_. _)" [0, 51, 10] 10) - translations \ \Beware of argument permutation!\ - "PROD x<-xs. b" == "CONST listprod (CONST map (%x. b) xs)" - "\x\xs. b" == "CONST listprod (CONST map (%x. b) xs)" + "\x\xs. b" \ "CONST listprod (CONST map (\x. b) xs)" end diff -r 1d43f86f48be -r e96292f32c3c src/HOL/HOL.thy --- a/src/HOL/HOL.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/HOL.thy Mon Dec 28 21:47:32 2015 +0100 @@ -73,48 +73,45 @@ Trueprop :: "bool \ prop" ("(_)" 5) axiomatization - implies :: "[bool, bool] \ bool" (infixr "-->" 25) and + implies :: "[bool, bool] \ bool" (infixr "\" 25) and eq :: "['a, 'a] \ bool" (infixl "=" 50) and The :: "('a \ bool) \ 'a" consts True :: bool False :: bool - Not :: "bool \ bool" ("~ _" [40] 40) + Not :: "bool \ bool" ("\ _" [40] 40) - conj :: "[bool, bool] \ bool" (infixr "&" 35) - disj :: "[bool, bool] \ bool" (infixr "|" 30) + conj :: "[bool, bool] \ bool" (infixr "\" 35) + disj :: "[bool, bool] \ bool" (infixr "\" 30) - All :: "('a \ bool) \ bool" (binder "ALL " 10) - Ex :: "('a \ bool) \ bool" (binder "EX " 10) - Ex1 :: "('a \ bool) \ bool" (binder "EX! " 10) + All :: "('a \ bool) \ bool" (binder "\" 10) + Ex :: "('a \ bool) \ bool" (binder "\" 10) + Ex1 :: "('a \ bool) \ bool" (binder "\!" 10) subsubsection \Additional concrete syntax\ -notation (output) - eq (infix "=" 50) - -abbreviation - not_equal :: "['a, 'a] \ bool" (infixl "~=" 50) where - "x ~= y \ ~ (x = y)" +abbreviation not_equal :: "['a, 'a] \ bool" (infixl "\" 50) + where "x \ y \ \ (x = y)" notation (output) + eq (infix "=" 50) and + not_equal (infix "\" 50) + +notation (ASCII output) not_equal (infix "~=" 50) -notation (xsymbols) - Not ("\ _" [40] 40) and - conj (infixr "\" 35) and - disj (infixr "\" 30) and - implies (infixr "\" 25) and - not_equal (infixl "\" 50) - -notation (xsymbols output) - not_equal (infix "\" 50) +notation (ASCII) + Not ("~ _" [40] 40) and + conj (infixr "&" 35) and + disj (infixr "|" 30) and + implies (infixr "-->" 25) and + not_equal (infixl "~=" 50) abbreviation (iff) - iff :: "[bool, bool] \ bool" (infixr "\" 25) where - "A \ B \ A = B" + iff :: "[bool, bool] \ bool" (infixr "\" 25) + where "A \ B \ A = B" syntax "_The" :: "[pttrn, bool] \ 'a" ("(3THE _./ _)" [0, 10] 10) translations "THE x. P" \ "CONST The (\x. P)" @@ -134,16 +131,16 @@ nonterminal case_syn and cases_syn syntax "_case_syntax" :: "['a, cases_syn] \ 'b" ("(case _ of/ _)" 10) - "_case1" :: "['a, 'b] \ case_syn" ("(2_ =>/ _)" 10) + "_case1" :: "['a, 'b] \ case_syn" ("(2_ \/ _)" 10) "" :: "case_syn \ cases_syn" ("_") "_case2" :: "[case_syn, cases_syn] \ cases_syn" ("_/ | _") -syntax (xsymbols) - "_case1" :: "['a, 'b] \ case_syn" ("(2_ \/ _)" 10) +syntax (ASCII) + "_case1" :: "['a, 'b] \ case_syn" ("(2_ =>/ _)" 10) -notation (xsymbols) - All (binder "\" 10) and - Ex (binder "\" 10) and - Ex1 (binder "\!" 10) +notation (ASCII) + All (binder "ALL " 10) and + Ex (binder "EX " 10) and + Ex1 (binder "EX! " 10) notation (HOL) All (binder "! " 10) and diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Inductive.thy Mon Dec 28 21:47:32 2015 +0100 @@ -370,13 +370,11 @@ ML_file "Tools/BNF/bnf_fp_rec_sugar_util.ML" ML_file "Tools/BNF/bnf_lfp_rec_sugar.ML" -text\Lambda-abstractions with pattern matching:\ - +text \Lambda-abstractions with pattern matching:\ +syntax (ASCII) + "_lam_pats_syntax" :: "cases_syn \ 'a \ 'b" ("(%_)" 10) syntax - "_lam_pats_syntax" :: "cases_syn => 'a => 'b" ("(%_)" 10) -syntax (xsymbols) - "_lam_pats_syntax" :: "cases_syn => 'a => 'b" ("(\_)" 10) - + "_lam_pats_syntax" :: "cases_syn \ 'a \ 'b" ("(\_)" 10) parse_translation \ let fun fun_tr ctxt [cs] = diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/FinFun.thy --- a/src/HOL/Library/FinFun.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/FinFun.thy Mon Dec 28 21:47:32 2015 +0100 @@ -907,10 +907,11 @@ subsection \Function composition\ -definition finfun_comp :: "('a \ 'b) \ 'c \f 'a \ 'c \f 'b" (infixr "o$" 55) -where [code del]: "g o$ f = finfun_rec (\b. (K$ g b)) (\a b c. c(a $:= g b)) f" +definition finfun_comp :: "('a \ 'b) \ 'c \f 'a \ 'c \f 'b" (infixr "\$" 55) +where [code del]: "g \$ f = finfun_rec (\b. (K$ g b)) (\a b c. c(a $:= g b)) f" -notation (xsymbols) finfun_comp (infixr "\$" 55) +notation (ASCII) + finfun_comp (infixr "o$" 55) interpretation finfun_comp_aux: finfun_rec_wf_aux "(\b. (K$ g b))" "(\a b c. c(a $:= g b))" by(unfold_locales)(auto simp add: finfun_upd_apply intro: finfun_ext) @@ -968,10 +969,11 @@ thus ?thesis by(auto simp add: fun_eq_iff) qed -definition finfun_comp2 :: "'b \f 'c \ ('a \ 'b) \ 'a \f 'c" (infixr "$o" 55) -where [code del]: "g $o f = Abs_finfun (op $ g \ f)" +definition finfun_comp2 :: "'b \f 'c \ ('a \ 'b) \ 'a \f 'c" (infixr "$\" 55) +where [code del]: "g $\ f = Abs_finfun (op $ g \ f)" -notation (xsymbol) finfun_comp2 (infixr "$\" 55) +notation (ASCII) + finfun_comp2 (infixr "$o" 55) lemma finfun_comp2_const [code, simp]: "finfun_comp2 (K$ c) f = (K$ c)" including finfun @@ -1534,12 +1536,12 @@ finfun_const ("K$/ _" [0] 1) and finfun_update ("_'(_ $:= _')" [1000,0,0] 1000) and finfun_apply (infixl "$" 999) and - finfun_comp (infixr "o$" 55) and - finfun_comp2 (infixr "$o" 55) and + finfun_comp (infixr "\$" 55) and + finfun_comp2 (infixr "$\" 55) and finfun_Diag ("(1'($_,/ _$'))" [0, 0] 1000) -no_notation (xsymbols) - finfun_comp (infixr "\$" 55) and - finfun_comp2 (infixr "$\" 55) +no_notation (ASCII) + finfun_comp (infixr "o$" 55) and + finfun_comp2 (infixr "$o" 55) end diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/FinFun_Syntax.thy --- a/src/HOL/Library/FinFun_Syntax.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/FinFun_Syntax.thy Mon Dec 28 21:47:32 2015 +0100 @@ -13,12 +13,12 @@ finfun_const ("K$/ _" [0] 1) and finfun_update ("_'(_ $:= _')" [1000,0,0] 1000) and finfun_apply (infixl "$" 999) and - finfun_comp (infixr "o$" 55) and - finfun_comp2 (infixr "$o" 55) and + finfun_comp (infixr "\$" 55) and + finfun_comp2 (infixr "$\" 55) and finfun_Diag ("(1'($_,/ _$'))" [0, 0] 1000) -notation (xsymbols) - finfun_comp (infixr "\$" 55) and - finfun_comp2 (infixr "$\" 55) +notation (ASCII) + finfun_comp (infixr "o$" 55) and + finfun_comp2 (infixr "$o" 55) end \ No newline at end of file diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/FuncSet.thy Mon Dec 28 21:47:32 2015 +0100 @@ -20,10 +20,10 @@ abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 60) where "A \ B \ Pi A (\_. B)" -syntax +syntax (ASCII) "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3PI _:_./ _)" 10) "_lam" :: "pttrn \ 'a set \ 'a \ 'b \ ('a \ 'b)" ("(3%_:_./ _)" [0,0,3] 3) -syntax (xsymbols) +syntax "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\ _\_./ _)" 10) "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" ("(3\_\_./ _)" [0,0,3] 3) translations @@ -347,11 +347,12 @@ abbreviation "Pi\<^sub>E A B \ PiE A B" -syntax +syntax (ASCII) "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3PIE _:_./ _)" 10) -syntax (xsymbols) +syntax "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\\<^sub>E _\_./ _)" 10) -translations "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" +translations + "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\\<^sub>E" 60) where "A \\<^sub>E B \ (\\<^sub>E i\A. B)" diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/Groups_Big_Fun.thy --- a/src/HOL/Library/Groups_Big_Fun.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/Groups_Big_Fun.thy Mon Dec 28 21:47:32 2015 +0100 @@ -233,12 +233,12 @@ end -syntax +syntax (ASCII) "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3SUM _. _)" [0, 10] 10) -syntax (xsymbols) +syntax "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3\_. _)" [0, 10] 10) translations - "\a. b" == "CONST Sum_any (\a. b)" + "\a. b" \ "CONST Sum_any (\a. b)" lemma Sum_any_left_distrib: fixes r :: "'a :: semiring_0" @@ -302,10 +302,10 @@ end +syntax (ASCII) + "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3PROD _. _)" [0, 10] 10) syntax - "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3PROD _. _)" [0, 10] 10) -syntax (xsymbols) - "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3\_. _)" [0, 10] 10) + "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3\_. _)" [0, 10] 10) translations "\a. b" == "CONST Prod_any (\a. b)" diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/Lattice_Syntax.thy --- a/src/HOL/Library/Lattice_Syntax.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/Lattice_Syntax.thy Mon Dec 28 21:47:32 2015 +0100 @@ -14,7 +14,7 @@ Inf ("\_" [900] 900) and Sup ("\_" [900] 900) -syntax (xsymbols) +syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/Monad_Syntax.thy --- a/src/HOL/Library/Monad_Syntax.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/Monad_Syntax.thy Mon Dec 28 21:47:32 2015 +0100 @@ -15,42 +15,42 @@ \ consts - bind :: "['a, 'b \ 'c] \ 'd" (infixr ">>=" 54) + bind :: "['a, 'b \ 'c] \ 'd" (infixr "\=" 54) -notation (xsymbols) - bind (infixr "\=" 54) +notation (ASCII) + bind (infixr ">>=" 54) notation (latex output) bind (infixr "\" 54) + abbreviation (do_notation) bind_do :: "['a, 'b \ 'c] \ 'd" -where - "bind_do \ bind" + where "bind_do \ bind" notation (output) - bind_do (infixr ">>=" 54) + bind_do (infixr "\=" 54) -notation (xsymbols output) - bind_do (infixr "\=" 54) +notation (ASCII output) + bind_do (infixr ">>=" 54) notation (latex output) bind_do (infixr "\" 54) + nonterminal do_binds and do_bind - syntax "_do_block" :: "do_binds \ 'a" ("do {//(2 _)//}" [12] 62) - "_do_bind" :: "[pttrn, 'a] \ do_bind" ("(2_ <-/ _)" 13) + "_do_bind" :: "[pttrn, 'a] \ do_bind" ("(2_ \/ _)" 13) "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" [1000, 13] 13) "_do_then" :: "'a \ do_bind" ("_" [14] 13) "_do_final" :: "'a \ do_binds" ("_") "_do_cons" :: "[do_bind, do_binds] \ do_binds" ("_;//_" [13, 12] 12) - "_thenM" :: "['a, 'b] \ 'c" (infixr ">>" 54) + "_thenM" :: "['a, 'b] \ 'c" (infixr "\" 54) -syntax (xsymbols) - "_do_bind" :: "[pttrn, 'a] \ do_bind" ("(2_ \/ _)" 13) - "_thenM" :: "['a, 'b] \ 'c" (infixr "\" 54) +syntax (ASCII) + "_do_bind" :: "[pttrn, 'a] \ do_bind" ("(2_ <-/ _)" 13) + "_thenM" :: "['a, 'b] \ 'c" (infixr ">>" 54) syntax (latex output) "_thenM" :: "['a, 'b] \ 'c" (infixr "\" 54) diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/Multiset.thy Mon Dec 28 21:47:32 2015 +0100 @@ -25,11 +25,11 @@ setup_lifting type_definition_multiset -abbreviation Melem :: "'a \ 'a multiset \ bool" ("(_/ :# _)" [50, 51] 50) where - "a :# M == 0 < count M a" - -notation (xsymbols) - Melem (infix "\#" 50) +abbreviation Melem :: "'a \ 'a multiset \ bool" (infix "\#" 50) + where "a \# M \ 0 < count M a" + +notation (ASCII) + Melem ("(_/ :# _)" [50, 51] 50) (* FIXME !? *) lemma multiset_eq_iff: "M = N \ (\a. count M a = count N a)" by (simp only: count_inject [symmetric] fun_eq_iff) @@ -265,16 +265,18 @@ subsubsection \Pointwise ordering induced by count\ -definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "<=#" 50) where -"subseteq_mset A B = (\a. count A a \ count B a)" - -definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix "<#" 50) where -"subset_mset A B = (A <=# B \ A \ B)" - -notation subseteq_mset (infix "\#" 50) -notation (xsymbols) subseteq_mset (infix "\#" 50) - -notation (xsymbols) subset_mset (infix "\#" 50) +definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) + where "A \# B = (\a. count A a \ count B a)" + +definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) + where "A \# B = (A \# B \ A \ B)" + +notation (input) + subseteq_mset (infix "\#" 50) + +notation (ASCII) + subseteq_mset (infix "<=#" 50) and + subset_mset (infix "<#" 50) interpretation subset_mset: ordered_ab_semigroup_add_imp_le "op +" "op -" "op \#" "op \#" by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym) @@ -504,9 +506,9 @@ show ?thesis unfolding B by auto qed -syntax +syntax (ASCII) "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{# _ :# _./ _#})") -syntax (xsymbol) +syntax "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{# _ \# _./ _#})") translations "{#x \# M. P#}" == "CONST filter_mset (\x. P) M" @@ -855,27 +857,17 @@ lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" by (cases M) auto +syntax (ASCII) + "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ :# _#})") syntax - "_comprehension1_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" - ("({#_/. _ :# _#})") -translations - "{#e. x:#M#}" == "CONST image_mset (\x. e) M" - -syntax (xsymbols) - "_comprehension2_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" - ("({#_/. _ \# _#})") + "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ \# _#})") translations - "{#e. x \# M#}" == "CONST image_mset (\x. e) M" - + "{#e. x \# M#}" \ "CONST image_mset (\x. e) M" + +syntax (ASCII) + "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ :# _./ _#})") syntax - "_comprehension3_mset" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" - ("({#_/ | _ :# _./ _#})") -translations - "{#e | x:#M. P#}" \ "{#e. x :# {# x:#M. P#}#}" - -syntax - "_comprehension4_mset" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" - ("({#_/ | _ \# _./ _#})") + "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ \# _./ _#})") translations "{#e | x\#M. P#}" \ "{#e. x \# {# x\#M. P#}#}" @@ -1235,10 +1227,8 @@ qed -abbreviation Union_mset :: "'a multiset multiset \ 'a multiset" where - "Union_mset MM \ msetsum MM" - -notation (xsymbols) Union_mset ("\#_" [900] 900) +abbreviation Union_mset :: "'a multiset multiset \ 'a multiset" ("\#_" [900] 900) + where "\# MM \ msetsum MM" lemma set_mset_Union_mset[simp]: "set_mset (\# MM) = (\M \ set_mset MM. set_mset M)" by (induct MM) auto @@ -1246,14 +1236,12 @@ lemma in_Union_mset_iff[iff]: "x \# \# MM \ (\M. M \# MM \ x \# M)" by (induct MM) auto +syntax (ASCII) + "_msetsum_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10) syntax - "_msetsum_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" - ("(3SUM _:#_. _)" [0, 51, 10] 10) -syntax (xsymbols) - "_msetsum_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" - ("(3\_\#_. _)" [0, 51, 10] 10) + "_msetsum_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3\_\#_. _)" [0, 51, 10] 10) translations - "SUM i :# A. b" == "CONST msetsum (CONST image_mset (\i. b) A)" + "\i \# A. b" \ "CONST msetsum (CONST image_mset (\i. b) A)" context comm_monoid_mult begin @@ -1287,14 +1275,12 @@ end +syntax (ASCII) + "_msetprod_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3PROD _:#_. _)" [0, 51, 10] 10) syntax - "_msetprod_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" - ("(3PROD _:#_. _)" [0, 51, 10] 10) -syntax (xsymbols) - "_msetprod_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" - ("(3\_\#_. _)" [0, 51, 10] 10) + "_msetprod_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3\_\#_. _)" [0, 51, 10] 10) translations - "PROD i :# A. b" == "CONST msetprod (CONST image_mset (\i. b) A)" + "\i \# A. b" \ "CONST msetprod (CONST image_mset (\i. b) A)" lemma (in comm_semiring_1) dvd_msetprod: assumes "x \# A" @@ -1718,14 +1704,15 @@ subsubsection \Partial-order properties\ -definition less_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#<#" 50) where - "M' #<# M \ (M', M) \ mult {(x', x). x' < x}" - -definition le_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#<=#" 50) where - "M' #<=# M \ M' #<# M \ M' = M" - -notation (xsymbols) less_multiset (infix "#\#" 50) -notation (xsymbols) le_multiset (infix "#\#" 50) +definition less_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#\#" 50) + where "M' #\# M \ (M', M) \ mult {(x', x). x' < x}" + +definition le_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#\#" 50) + where "M' #\# M \ M' #\# M \ M' = M" + +notation (ASCII) + less_multiset (infix "#<#" 50) and + le_multiset (infix "#<=#" 50) interpretation multiset_order: order le_multiset less_multiset proof - diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/Option_ord.thy --- a/src/HOL/Library/Option_ord.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/Option_ord.thy Mon Dec 28 21:47:32 2015 +0100 @@ -16,7 +16,7 @@ Inf ("\_" [900] 900) and Sup ("\_" [900] 900) -syntax (xsymbols) +syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) @@ -348,7 +348,7 @@ Inf ("\_" [900] 900) and Sup ("\_" [900] 900) -no_syntax (xsymbols) +no_syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/OptionalSugar.thy --- a/src/HOL/Library/OptionalSugar.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/OptionalSugar.thy Mon Dec 28 21:47:32 2015 +0100 @@ -63,7 +63,7 @@ (* sorts as intersections *) -syntax (xsymbols output) +syntax (output) "_topsort" :: "sort" ("\" 1000) "_sort" :: "classes => sort" ("'(_')" 1000) "_classes" :: "id => classes => classes" ("_ \ _" 1000) diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Library/State_Monad.thy --- a/src/HOL/Library/State_Monad.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Library/State_Monad.thy Mon Dec 28 21:47:32 2015 +0100 @@ -118,14 +118,14 @@ syntax "_sdo_block" :: "sdo_binds \ 'a" ("exec {//(2 _)//}" [12] 62) - "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" ("(_ <-/ _)" 13) + "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" ("(_ \/ _)" 13) "_sdo_let" :: "[pttrn, 'a] \ sdo_bind" ("(2let _ =/ _)" [1000, 13] 13) "_sdo_then" :: "'a \ sdo_bind" ("_" [14] 13) "_sdo_final" :: "'a \ sdo_binds" ("_") "_sdo_cons" :: "[sdo_bind, sdo_binds] \ sdo_binds" ("_;//_" [13, 12] 12) -syntax (xsymbols) - "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" ("(_ \/ _)" 13) +syntax (ASCII) + "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" ("(_ <-/ _)" 13) translations "_sdo_block (_sdo_cons (_sdo_bind p t) (_sdo_final e))" diff -r 1d43f86f48be -r e96292f32c3c src/HOL/List.thy --- a/src/HOL/List.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/List.thy Mon Dec 28 21:47:32 2015 +0100 @@ -77,13 +77,13 @@ "filter P [] = []" | "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)" +text \Special syntax for filter:\ +syntax (ASCII) + "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])") syntax - \ \Special syntax for filter\ - "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])") -syntax (xsymbols) - "_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\_ ./ _])") + "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_\_ ./ _])") translations - "[x<-xs . P]"== "CONST filter (%x. P) xs" + "[x<-xs . P]" \ "CONST filter (\x. P) xs" primrec fold :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where fold_Nil: "fold f [] = id" | @@ -395,13 +395,16 @@ syntax "_listcompr" :: "'a \ lc_qual \ lc_quals \ 'a list" ("[_ . __") - "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ <- _") + "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ \ _") "_lc_test" :: "bool \ lc_qual" ("_") (*"_lc_let" :: "letbinds => lc_qual" ("let _")*) "_lc_end" :: "lc_quals" ("]") "_lc_quals" :: "lc_qual \ lc_quals \ lc_quals" (", __") "_lc_abs" :: "'a => 'b list => 'b list" +syntax (ASCII) + "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ <- _") + (* These are easier than ML code but cannot express the optimized translation of [e. p<-xs] translations @@ -415,9 +418,6 @@ => "_Let b (_listcompr e Q Qs)" *) -syntax (xsymbols) - "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ \ _") - parse_translation \ let val NilC = Syntax.const @{const_syntax Nil}; diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Main.thy --- a/src/HOL/Main.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Main.thy Mon Dec 28 21:47:32 2015 +0100 @@ -35,7 +35,7 @@ hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV -no_syntax (xsymbols) +no_syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Map.thy --- a/src/HOL/Map.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Map.thy Mon Dec 28 21:47:32 2015 +0100 @@ -47,16 +47,16 @@ nonterminal maplets and maplet syntax - "_maplet" :: "['a, 'a] \ maplet" ("_ /|->/ _") - "_maplets" :: "['a, 'a] \ maplet" ("_ /[|->]/ _") + "_maplet" :: "['a, 'a] \ maplet" ("_ /\/ _") + "_maplets" :: "['a, 'a] \ maplet" ("_ /[\]/ _") "" :: "maplet \ maplets" ("_") "_Maplets" :: "[maplet, maplets] \ maplets" ("_,/ _") - "_MapUpd" :: "['a \ 'b, maplets] \ 'a \ 'b" ("_/'(_')" [900,0]900) + "_MapUpd" :: "['a \ 'b, maplets] \ 'a \ 'b" ("_/'(_')" [900, 0] 900) "_Map" :: "maplets \ 'a \ 'b" ("(1[_])") -syntax (xsymbols) - "_maplet" :: "['a, 'a] \ maplet" ("_ /\/ _") - "_maplets" :: "['a, 'a] \ maplet" ("_ /[\]/ _") +syntax (ASCII) + "_maplet" :: "['a, 'a] \ maplet" ("_ /|->/ _") + "_maplets" :: "['a, 'a] \ maplet" ("_ /[|->]/ _") translations "_MapUpd m (_Maplets xy ms)" \ "_MapUpd (_MapUpd m xy) ms" @@ -65,15 +65,13 @@ "_Map (_Maplets ms1 ms2)" \ "_MapUpd (_Map ms1) ms2" "_Maplets ms1 (_Maplets ms2 ms3)" \ "_Maplets (_Maplets ms1 ms2) ms3" -primrec - map_of :: "('a \ 'b) list \ 'a \ 'b" where - "map_of [] = empty" - | "map_of (p # ps) = (map_of ps)(fst p \ snd p)" +primrec map_of :: "('a \ 'b) list \ 'a \ 'b" +where + "map_of [] = empty" +| "map_of (p # ps) = (map_of ps)(fst p \ snd p)" -definition - map_upds :: "('a \ 'b) \ 'a list \ 'b list \ 'a \ 'b" where - "map_upds m xs ys = m ++ map_of (rev (zip xs ys))" - +definition map_upds :: "('a \ 'b) \ 'a list \ 'b list \ 'a \ 'b" + where "map_upds m xs ys = m ++ map_of (rev (zip xs ys))" translations "_MapUpd m (_maplets x y)" \ "CONST map_upds m x y" diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Orderings.thy Mon Dec 28 21:47:32 2015 +0100 @@ -93,25 +93,25 @@ begin notation - less_eq ("op <=") and - less_eq ("(_/ <= _)" [51, 51] 50) and + less_eq ("op \") and + less_eq ("(_/ \ _)" [51, 51] 50) and less ("op <") and less ("(_/ < _)" [51, 51] 50) -notation (xsymbols) - less_eq ("op \") and - less_eq ("(_/ \ _)" [51, 51] 50) +abbreviation (input) + greater_eq (infix "\" 50) + where "x \ y \ y \ x" abbreviation (input) - greater_eq (infix ">=" 50) where - "x >= y \ y <= x" + greater (infix ">" 50) + where "x > y \ y < x" + +notation (ASCII) + less_eq ("op <=") and + less_eq ("(_/ <= _)" [51, 51] 50) notation (input) - greater_eq (infix "\" 50) - -abbreviation (input) - greater (infix ">" 50) where - "x > y \ y < x" + greater_eq (infix ">=" 50) end @@ -652,7 +652,7 @@ subsection \Bounded quantifiers\ -syntax +syntax (ASCII) "_All_less" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) @@ -663,7 +663,7 @@ "_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _>=_./ _)" [0, 0, 10] 10) "_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3EX _>=_./ _)" [0, 0, 10] 10) -syntax (xsymbols) +syntax "_All_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Power.thy --- a/src/HOL/Power.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Power.thy Mon Dec 28 21:47:32 2015 +0100 @@ -27,21 +27,17 @@ class power = one + times begin -primrec power :: "'a \ nat \ 'a" (infixr "^" 80) where - power_0: "a ^ 0 = 1" - | power_Suc: "a ^ Suc n = a * a ^ n" +primrec power :: "'a \ nat \ 'a" (infixr "^" 80) +where + power_0: "a ^ 0 = 1" +| power_Suc: "a ^ Suc n = a * a ^ n" notation (latex output) power ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \Special syntax for squares.\ - -abbreviation (xsymbols) - power2 :: "'a \ 'a" ("(_\<^sup>2)" [1000] 999) where - "x\<^sup>2 \ x ^ 2" - -notation (latex output) - power2 ("(_\<^sup>2)" [1000] 999) +abbreviation power2 :: "'a \ 'a" ("(_\<^sup>2)" [1000] 999) + where "x\<^sup>2 \ x ^ 2" end diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Product_Type.thy Mon Dec 28 21:47:32 2015 +0100 @@ -226,11 +226,11 @@ definition "prod = {f. \a b. f = Pair_Rep (a::'a) (b::'b)}" -typedef ('a, 'b) prod (infixr "*" 20) = "prod :: ('a \ 'b \ bool) set" +typedef ('a, 'b) prod ("(_ \/ _)" [21, 20] 20) = "prod :: ('a \ 'b \ bool) set" unfolding prod_def by auto -type_notation (xsymbols) - "prod" ("(_ \/ _)" [21, 20] 20) +type_notation (ASCII) + prod (infixr "*" 20) definition Pair :: "'a \ 'b \ 'a \ 'b" where "Pair a b = Abs_prod (Pair_Rep a b)" diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Record.thy --- a/src/HOL/Record.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Record.thy Mon Dec 28 21:47:32 2015 +0100 @@ -427,26 +427,26 @@ "_field_type" :: "ident => type => field_type" ("(2_ ::/ _)") "" :: "field_type => field_types" ("_") "_field_types" :: "field_type => field_types => field_types" ("_,/ _") - "_record_type" :: "field_types => type" ("(3'(| _ |'))") - "_record_type_scheme" :: "field_types => type => type" ("(3'(| _,/ (2... ::/ _) |'))") + "_record_type" :: "field_types => type" ("(3\_\)") + "_record_type_scheme" :: "field_types => type => type" ("(3\_,/ (2\ ::/ _)\)") "_field" :: "ident => 'a => field" ("(2_ =/ _)") "" :: "field => fields" ("_") "_fields" :: "field => fields => fields" ("_,/ _") - "_record" :: "fields => 'a" ("(3'(| _ |'))") - "_record_scheme" :: "fields => 'a => 'a" ("(3'(| _,/ (2... =/ _) |'))") + "_record" :: "fields => 'a" ("(3\_\)") + "_record_scheme" :: "fields => 'a => 'a" ("(3\_,/ (2\ =/ _)\)") "_field_update" :: "ident => 'a => field_update" ("(2_ :=/ _)") "" :: "field_update => field_updates" ("_") "_field_updates" :: "field_update => field_updates => field_updates" ("_,/ _") - "_record_update" :: "'a => field_updates => 'b" ("_/(3'(| _ |'))" [900, 0] 900) + "_record_update" :: "'a => field_updates => 'b" ("_/(3\_\)" [900, 0] 900) -syntax (xsymbols) - "_record_type" :: "field_types => type" ("(3\_\)") - "_record_type_scheme" :: "field_types => type => type" ("(3\_,/ (2\ ::/ _)\)") - "_record" :: "fields => 'a" ("(3\_\)") - "_record_scheme" :: "fields => 'a => 'a" ("(3\_,/ (2\ =/ _)\)") - "_record_update" :: "'a => field_updates => 'b" ("_/(3\_\)" [900, 0] 900) +syntax (ASCII) + "_record_type" :: "field_types => type" ("(3'(| _ |'))") + "_record_type_scheme" :: "field_types => type => type" ("(3'(| _,/ (2... ::/ _) |'))") + "_record" :: "fields => 'a" ("(3'(| _ |'))") + "_record_scheme" :: "fields => 'a => 'a" ("(3'(| _,/ (2... =/ _) |'))") + "_record_update" :: "'a => field_updates => 'b" ("_/(3'(| _ |'))" [900, 0] 900) subsection \Record package\ diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Relation.thy --- a/src/HOL/Relation.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Relation.thy Mon Dec 28 21:47:32 2015 +0100 @@ -684,19 +684,17 @@ subsubsection \Converse\ -inductive_set converse :: "('a \ 'b) set \ ('b \ 'a) set" ("(_^-1)" [1000] 999) +inductive_set converse :: "('a \ 'b) set \ ('b \ 'a) set" ("(_\)" [1000] 999) for r :: "('a \ 'b) set" where - "(a, b) \ r \ (b, a) \ r^-1" - -notation (xsymbols) - converse ("(_\)" [1000] 999) + "(a, b) \ r \ (b, a) \ r\" notation - conversep ("(_^--1)" [1000] 1000) + conversep ("(_\\)" [1000] 1000) -notation (xsymbols) - conversep ("(_\\)" [1000] 1000) +notation (ASCII) + converse ("(_^-1)" [1000] 999) and + conversep ("(_^--1)" [1000] 1000) lemma converseI [sym]: "(a, b) \ r \ (b, a) \ r\" diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Set.thy --- a/src/HOL/Set.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Set.thy Mon Dec 28 21:47:32 2015 +0100 @@ -12,41 +12,39 @@ axiomatization Collect :: "('a \ bool) \ 'a set" \ "comprehension" and member :: "'a \ 'a set \ bool" \ "membership" -where - mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a" +where mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a" and Collect_mem_eq [simp]: "Collect (\x. member x A) = A" notation + member ("op \") and + member ("(_/ \ _)" [51, 51] 50) + +abbreviation not_member + where "not_member x A \ \ (x \ A)" \ "non-membership" +notation + not_member ("op \") and + not_member ("(_/ \ _)" [51, 51] 50) + +notation (ASCII) member ("op :") and - member ("(_/ : _)" [51, 51] 50) - -abbreviation not_member where - "not_member x A \ ~ (x : A)" \ "non-membership" - -notation + member ("(_/ : _)" [51, 51] 50) and not_member ("op ~:") and not_member ("(_/ ~: _)" [51, 51] 50) -notation (xsymbols) - member ("op \") and - member ("(_/ \ _)" [51, 51] 50) and - not_member ("op \") and - not_member ("(_/ \ _)" [51, 51] 50) - text \Set comprehensions\ syntax "_Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") translations - "{x. P}" == "CONST Collect (%x. P)" - + "{x. P}" \ "CONST Collect (\x. P)" + +syntax (ASCII) + "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{_ :/ _./ _})") syntax - "_Collect" :: "pttrn => 'a set => bool => 'a set" ("(1{_ :/ _./ _})") -syntax (xsymbols) - "_Collect" :: "pttrn => 'a set => bool => 'a set" ("(1{_ \/ _./ _})") + "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{_ \/ _./ _})") translations - "{p:A. P}" => "CONST Collect (%p. p:A & P)" + "{p:A. P}" \ "CONST Collect (\p. p \ A \ P)" lemma CollectI: "P a \ a \ {x. P x}" by simp @@ -141,21 +139,13 @@ subsection \Subsets and bounded quantifiers\ -abbreviation - subset :: "'a set \ 'a set \ bool" where - "subset \ less" - -abbreviation - subset_eq :: "'a set \ 'a set \ bool" where - "subset_eq \ less_eq" - -notation (output) - subset ("op <") and - subset ("(_/ < _)" [51, 51] 50) and - subset_eq ("op <=") and - subset_eq ("(_/ <= _)" [51, 51] 50) - -notation (xsymbols) +abbreviation subset :: "'a set \ 'a set \ bool" + where "subset \ less" + +abbreviation subset_eq :: "'a set \ 'a set \ bool" + where "subset_eq \ less_eq" + +notation subset ("op \") and subset ("(_/ \ _)" [51, 51] 50) and subset_eq ("op \") and @@ -169,19 +159,25 @@ supset_eq :: "'a set \ 'a set \ bool" where "supset_eq \ greater_eq" -notation (xsymbols) +notation supset ("op \") and supset ("(_/ \ _)" [51, 51] 50) and supset_eq ("op \") and supset_eq ("(_/ \ _)" [51, 51] 50) +notation (ASCII output) + subset ("op <") and + subset ("(_/ < _)" [51, 51] 50) and + subset_eq ("op <=") and + subset_eq ("(_/ <= _)" [51, 51] 50) + definition Ball :: "'a set \ ('a \ bool) \ bool" where "Ball A P \ (\x. x \ A \ P x)" \ "bounded universal quantifiers" definition Bex :: "'a set \ ('a \ bool) \ bool" where "Bex A P \ (\x. x \ A \ P x)" \ "bounded existential quantifiers" -syntax +syntax (ASCII) "_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3EX! _:_./ _)" [0, 0, 10] 10) @@ -192,32 +188,25 @@ "_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3?! _:_./ _)" [0, 0, 10] 10) -syntax (xsymbols) +syntax "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST_\_./ _)" [0, 0, 10] 10) translations - "ALL x:A. P" == "CONST Ball A (%x. P)" - "EX x:A. P" == "CONST Bex A (%x. P)" - "EX! x:A. P" => "EX! x. x:A & P" - "LEAST x:A. P" => "LEAST x. x:A & P" - -syntax (output) + "\x\A. P" \ "CONST Ball A (\x. P)" + "\x\A. P" \ "CONST Bex A (\x. P)" + "\!x\A. P" \ "\!x. x \ A \ P" + "LEAST x:A. P" \ "LEAST x. x \ A \ P" + +syntax (ASCII output) "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_setleAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) "_setleEx" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3EX! _<=_./ _)" [0, 0, 10] 10) -syntax (xsymbols) - "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setleAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setleEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) - syntax (HOL output) "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10) @@ -225,12 +214,19 @@ "_setleEx" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3?! _<=_./ _)" [0, 0, 10] 10) +syntax + "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setleAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setleEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) + translations - "\A\B. P" => "ALL A. A \ B --> P" - "\A\B. P" => "EX A. A \ B & P" - "\A\B. P" => "ALL A. A \ B --> P" - "\A\B. P" => "EX A. A \ B & P" - "\!A\B. P" => "EX! A. A \ B & P" + "\A\B. P" \ "\A. A \ B \ P" + "\A\B. P" \ "\A. A \ B \ P" + "\A\B. P" \ "\A. A \ B \ P" + "\A\B. P" \ "\A. A \ B \ P" + "\!A\B. P" \ "\!A. A \ B \ P" print_translation \ let @@ -669,11 +665,11 @@ subsubsection \Binary intersection\ -abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where - "op Int \ inf" - -notation (xsymbols) - inter (infixl "\" 70) +abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "\" 70) + where "op \ \ inf" + +notation (ASCII) + inter (infixl "Int" 70) lemma Int_def: "A \ B = {x. x \ A \ x \ B}" @@ -700,11 +696,11 @@ subsubsection \Binary union\ -abbreviation union :: "'a set \ 'a set \ 'a set" (infixl "Un" 65) where - "union \ sup" - -notation (xsymbols) - union (infixl "\" 65) +abbreviation union :: "'a set \ 'a set \ 'a set" (infixl "\" 65) + where "union \ sup" + +notation (ASCII) + union (infixl "Un" 65) lemma Un_def: "A \ B = {x. x \ A \ x \ B}" diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Set_Interval.thy Mon Dec 28 21:47:32 2015 +0100 @@ -59,29 +59,29 @@ nat}: it is equivalent to @{term"{0::nat.. -syntax +syntax (ASCII) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) -syntax (xsymbols) - "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) - "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) - "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) - "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) - syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ < _)/ _)" [0, 0, 10] 10) +syntax + "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) + "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) + "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) + "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) + translations - "UN i<=n. A" == "UN i:{..n}. A" - "UN ii\n. A" \ "\i\{..n}. A" + "\i "\i\{..i\n. A" \ "\i\{..n}. A" + "\i "\i\{..Various equivalences\ @@ -1427,16 +1427,12 @@ subsection \Summation indexed over intervals\ -syntax - "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) - "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) - "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) - "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) -syntax (xsymbols) - "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) - "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) - "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) - "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) +syntax (ASCII) + "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) + "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) + "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) + "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) + syntax (latex_sum output) "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10) @@ -1447,9 +1443,15 @@ "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^raw:$\sum_{>_ \ _\<^raw:}$> _)" [0,0,10] 10) +syntax + "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) + "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) + "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) + "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) + translations - "\x=a..b. t" == "CONST setsum (%x. t) {a..b}" - "\x=a..x=a..b. t" == "CONST setsum (\x. t) {a..b}" + "\x=a..x. t) {a..i\n. t" == "CONST setsum (\i. t) {..n}" "\ii. t) {..x. Q x \ P x \ (\xxxProducts indexed over intervals\ -syntax - "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) - "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) - "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) - "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) -syntax (xsymbols) - "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) - "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) - "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) - "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) +syntax (ASCII) + "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) + "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) + "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) + "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) + syntax (latex_prod output) "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10) @@ -1805,11 +1804,18 @@ "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^raw:$\prod_{>_ \ _\<^raw:}$> _)" [0,0,10] 10) +syntax + "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) + "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) + "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) + "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) + translations - "\x=a..b. t" == "CONST setprod (%x. t) {a..b}" - "\x=a..i\n. t" == "CONST setprod (\i. t) {..n}" - "\ii. t) {..x=a..b. t" \ "CONST setprod (\x. t) {a..b}" + "\x=a.. "CONST setprod (\x. t) {a..i\n. t" \ "CONST setprod (\i. t) {..n}" + "\i "CONST setprod (\i. t) {..Transfer setup\ diff -r 1d43f86f48be -r e96292f32c3c src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Mon Dec 28 19:23:15 2015 +0100 +++ b/src/HOL/Transitive_Closure.thy Mon Dec 28 21:47:32 2015 +0100 @@ -24,46 +24,43 @@ notes [[inductive_defs]] begin -inductive_set - rtrancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_^*)" [1000] 999) +inductive_set rtrancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>*)" [1000] 999) + for r :: "('a \ 'a) set" +where + rtrancl_refl [intro!, Pure.intro!, simp]: "(a, a) \ r\<^sup>*" +| rtrancl_into_rtrancl [Pure.intro]: "(a, b) \ r\<^sup>* \ (b, c) \ r \ (a, c) \ r\<^sup>*" + +inductive_set trancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>+)" [1000] 999) for r :: "('a \ 'a) set" where - rtrancl_refl [intro!, Pure.intro!, simp]: "(a, a) : r^*" - | rtrancl_into_rtrancl [Pure.intro]: "(a, b) : r^* ==> (b, c) : r ==> (a, c) : r^*" + r_into_trancl [intro, Pure.intro]: "(a, b) \ r \ (a, b) \ r\<^sup>+" +| trancl_into_trancl [Pure.intro]: "(a, b) \ r\<^sup>+ \ (b, c) \ r \ (a, c) \ r\<^sup>+" -inductive_set - trancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_^+)" [1000] 999) - for r :: "('a \ 'a) set" -where - r_into_trancl [intro, Pure.intro]: "(a, b) : r ==> (a, b) : r^+" - | trancl_into_trancl [Pure.intro]: "(a, b) : r^+ ==> (b, c) : r ==> (a, c) : r^+" +notation + rtranclp ("(_\<^sup>*\<^sup>*)" [1000] 1000) and + tranclp ("(_\<^sup>+\<^sup>+)" [1000] 1000) -declare rtrancl_def [nitpick_unfold del] - rtranclp_def [nitpick_unfold del] - trancl_def [nitpick_unfold del] - tranclp_def [nitpick_unfold del] +declare + rtrancl_def [nitpick_unfold del] + rtranclp_def [nitpick_unfold del] + trancl_def [nitpick_unfold del] + tranclp_def [nitpick_unfold del] end -notation - rtranclp ("(_^**)" [1000] 1000) and - tranclp ("(_^++)" [1000] 1000) +abbreviation reflcl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>=)" [1000] 999) + where "r\<^sup>= \ r \ Id" -abbreviation - reflclp :: "('a => 'a => bool) => 'a => 'a => bool" ("(_^==)" [1000] 1000) where - "r^== \ sup r op =" +abbreviation reflclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" ("(_\<^sup>=\<^sup>=)" [1000] 1000) + where "r\<^sup>=\<^sup>= \ sup r op =" -abbreviation - reflcl :: "('a \ 'a) set => ('a \ 'a) set" ("(_^=)" [1000] 999) where - "r^= \ r \ Id" - -notation (xsymbols) - rtranclp ("(_\<^sup>*\<^sup>*)" [1000] 1000) and - tranclp ("(_\<^sup>+\<^sup>+)" [1000] 1000) and - reflclp ("(_\<^sup>=\<^sup>=)" [1000] 1000) and - rtrancl ("(_\<^sup>*)" [1000] 999) and - trancl ("(_\<^sup>+)" [1000] 999) and - reflcl ("(_\<^sup>=)" [1000] 999) +notation (ASCII) + rtrancl ("(_^*)" [1000] 999) and + trancl ("(_^+)" [1000] 999) and + reflcl ("(_^=)" [1000] 999) and + rtranclp ("(_^**)" [1000] 1000) and + tranclp ("(_^++)" [1000] 1000) and + reflclp ("(_^==)" [1000] 1000) subsection \Reflexive closure\