# HG changeset patch # User Christian Urban # Date 1260526445 -3600 # Node ID 587e97d144a0cbbb7ffc5e6bf8ab721709aeb8bf # Parent 6decb8811d30b6b211f6b5e67bf12d74fd2cd7cf# Parent f51c6069cd17fafc57b8a54b702a5c7ff81beb7d merged diff -r 6decb8811d30 -r 587e97d144a0 Quot/Examples/FSet.thy --- a/Quot/Examples/FSet.thy Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/Examples/FSet.thy Fri Dec 11 11:14:05 2009 +0100 @@ -27,18 +27,18 @@ done quotient_def - EMPTY :: "EMPTY :: 'a fset" -where + "EMPTY :: 'a fset" +as "[]::'a list" quotient_def - INSERT :: "INSERT :: 'a \ 'a fset \ 'a fset" -where + "INSERT :: 'a \ 'a fset \ 'a fset" +as "op #" quotient_def - FUNION :: "FUNION :: 'a fset \ 'a fset \ 'a fset" -where + "FUNION :: 'a fset \ 'a fset \ 'a fset" +as "op @" fun @@ -48,8 +48,8 @@ | card1_cons: "(card1 (x # xs)) = (if (x mem xs) then (card1 xs) else (Suc (card1 xs)))" quotient_def - CARD :: "CARD :: 'a fset \ nat" -where + "CARD :: 'a fset \ nat" +as "card1" (* text {* @@ -111,18 +111,18 @@ done quotient_def - IN :: "IN :: 'a \ 'a fset \ bool" -where + "IN :: 'a \ 'a fset \ bool" +as "op mem" quotient_def - FOLD :: "FOLD :: ('a \ 'a \ 'a) \ ('b \ 'a) \ 'a \ 'b fset \ 'a" -where + "FOLD :: ('a \ 'a \ 'a) \ ('b \ 'a) \ 'a \ 'b fset \ 'a" +as "fold1" quotient_def - fmap::"fmap :: ('a \ 'b) \ 'a fset \ 'b fset" -where + "fmap :: ('a \ 'b) \ 'a fset \ 'b fset" +as "map" lemma mem_rsp: @@ -275,13 +275,13 @@ by (rule equivp_list_eq) quotient_def - EMPTY2 :: "EMPTY2 :: 'a fset2" -where + "EMPTY2 :: 'a fset2" +as "[]::'a list" quotient_def - INSERT2 :: "INSERT2 :: 'a \ 'a fset2 \ 'a fset2" -where + "INSERT2 :: 'a \ 'a fset2 \ 'a fset2" +as "op #" lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" @@ -293,13 +293,13 @@ done quotient_def - fset_rec::"fset_rec :: 'a \ ('b \ 'b fset \ 'a \ 'a) \ 'b fset \ 'a" -where + "fset_rec :: 'a \ ('b \ 'b fset \ 'a \ 'a) \ 'b fset \ 'a" +as "list_rec" quotient_def - fset_case::"fset_case :: 'a \ ('b \ 'b fset \ 'a) \ 'b fset \ 'a" -where + "fset_case :: 'a \ ('b \ 'b fset \ 'a) \ 'b fset \ 'a" +as "list_case" (* Probably not true without additional assumptions about the function *) diff -r 6decb8811d30 -r 587e97d144a0 Quot/Examples/FSet2.thy --- a/Quot/Examples/FSet2.thy Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/Examples/FSet2.thy Fri Dec 11 11:14:05 2009 +0100 @@ -24,23 +24,23 @@ quotient fset = "'a list" / "list_eq" by (rule equivp_list_eq) -quotient_def - fempty :: "fempty :: 'a fset" ("{||}") -where +quotient_def + "fempty :: 'a fset" ("{||}") +as "[]" -quotient_def - finsert :: "finsert :: 'a \ 'a fset \ 'a fset" -where +quotient_def + "finsert :: 'a \ 'a fset \ 'a fset" +as "(op #)" lemma finsert_rsp[quot_respect]: shows "(op = ===> op \ ===> op \) (op #) (op #)" by (auto intro: list_eq.intros) -quotient_def - funion :: "funion :: 'a fset \ 'a fset \ 'a fset" ("_ \f _" [65,66] 65) -where +quotient_def + "funion :: 'a fset \ 'a fset \ 'a fset" ("_ \f _" [65,66] 65) +as "(op @)" lemma append_rsp_aux1: @@ -65,9 +65,9 @@ by (auto simp add: append_rsp_aux2) -quotient_def - fmem :: " fmem :: 'a \ 'a fset \ bool" ("_ \f _" [50, 51] 50) -where +quotient_def + "fmem :: 'a \ 'a fset \ bool" ("_ \f _" [50, 51] 50) +as "(op mem)" lemma memb_rsp_aux: @@ -89,9 +89,9 @@ where "inter_list X Y \ [x \ X. x\set Y]" -quotient_def - finter :: "finter::'a fset \ 'a fset \ 'a fset" ("_ \f _" [70, 71] 70) -where +quotient_def + "finter::'a fset \ 'a fset \ 'a fset" ("_ \f _" [70, 71] 70) +as "inter_list" no_syntax diff -r 6decb8811d30 -r 587e97d144a0 Quot/Examples/IntEx.thy --- a/Quot/Examples/IntEx.thy Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/Examples/IntEx.thy Fri Dec 11 11:14:05 2009 +0100 @@ -22,13 +22,13 @@ print_quotients quotient_def - ZERO::"zero :: my_int" -where + "ZERO :: my_int" +as "(0::nat, 0::nat)" quotient_def - ONE::"one :: my_int" -where + "ONE :: my_int" +as "(1::nat, 0::nat)" fun @@ -37,8 +37,8 @@ "my_plus (x, y) (u, v) = (x + u, y + v)" quotient_def - PLUS::"PLUS :: my_int \ my_int \ my_int" -where + "PLUS :: my_int \ my_int \ my_int" +as "my_plus" fun @@ -47,8 +47,8 @@ "my_neg (x, y) = (y, x)" quotient_def - NEG::"NEG :: my_int \ my_int" -where + "NEG :: my_int \ my_int" +as "my_neg" definition @@ -62,8 +62,8 @@ "my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)" quotient_def - MULT::"MULT :: my_int \ my_int \ my_int" -where + "MULT :: my_int \ my_int \ my_int" +as "my_mult" @@ -74,8 +74,8 @@ "my_le (x, y) (u, v) = (x+v \ u+y)" quotient_def - LE :: "LE :: my_int \ my_int \ bool" -where + "LE :: my_int \ my_int \ bool" +as "my_le" term LE diff -r 6decb8811d30 -r 587e97d144a0 Quot/Examples/IntEx2.thy --- a/Quot/Examples/IntEx2.thy Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/Examples/IntEx2.thy Fri Dec 11 11:14:05 2009 +0100 @@ -18,15 +18,13 @@ instantiation int :: "{zero, one, plus, minus, uminus, times, ord, abs, sgn}" begin -quotient_def - zero_int::"0 :: int" -where - "(0::nat, 0::nat)" +ML {* @{term "0 :: int"} *} quotient_def - one_int::"1 :: int" -where - "(1::nat, 0::nat)" + "0 :: int" as "(0::nat, 0::nat)" + +quotient_def + "1 :: int" as "(1::nat, 0::nat)" fun plus_raw :: "(nat \ nat) \ (nat \ nat) \ (nat \ nat)" @@ -34,9 +32,7 @@ "plus_raw (x, y) (u, v) = (x + u, y + v)" quotient_def - plus_int::"(op +) :: (int \ int \ int)" -where - "plus_raw" + "(op +) :: (int \ int \ int)" as "plus_raw" fun uminus_raw :: "(nat \ nat) \ (nat \ nat)" @@ -44,9 +40,7 @@ "uminus_raw (x, y) = (y, x)" quotient_def - uminus_int::"(uminus :: (int \ int))" -where - "uminus_raw" + "(uminus :: (int \ int))" as "uminus_raw" definition minus_int_def [code del]: "z - w = z + (-w::int)" @@ -57,9 +51,7 @@ "times_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)" quotient_def - times_int::"(op *) :: (int \ int \ int)" -where - "times_raw" + "(op *) :: (int \ int \ int)" as "times_raw" (* PROBLEM: this should be called le_int and le_raw / or odd *) fun @@ -68,9 +60,7 @@ "less_eq_raw (x, y) (u, v) = (x+v \ u+y)" quotient_def - less_eq_int :: "(op \) :: int \ int \ bool" -where - "less_eq_raw" + le_int_def: "(op \) :: int \ int \ bool" as "less_eq_raw" definition less_int_def [code del]: "(z\int) < w = (z \ w \ z \ w)" @@ -218,8 +208,7 @@ "int_of_nat_raw m = (m :: nat, 0 :: nat)" quotient_def - int_of_nat :: "int_of_nat :: nat \ int" -where "int_of_nat_raw" + "int_of_nat :: nat \ int" as "int_of_nat_raw" lemma[quot_respect]: shows "(op = ===> op \) int_of_nat_raw int_of_nat_raw" diff -r 6decb8811d30 -r 587e97d144a0 Quot/Examples/LFex.thy --- a/Quot/Examples/LFex.thy Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/Examples/LFex.thy Fri Dec 11 11:14:05 2009 +0100 @@ -81,49 +81,49 @@ by (auto intro: alpha_equivps) quotient_def - TYP :: "TYP :: KIND" -where + "TYP :: KIND" +as "Type" quotient_def - KPI :: "KPI :: TY \ name \ KIND \ KIND" -where + "KPI :: TY \ name \ KIND \ KIND" +as "KPi" quotient_def - TCONST :: "TCONST :: ident \ TY" -where + "TCONST :: ident \ TY" +as "TConst" quotient_def - TAPP :: "TAPP :: TY \ TRM \ TY" -where + "TAPP :: TY \ TRM \ TY" +as "TApp" quotient_def - TPI :: "TPI :: TY \ name \ TY \ TY" -where + "TPI :: TY \ name \ TY \ TY" +as "TPi" (* FIXME: does not work with CONST *) quotient_def - CONS :: "CONS :: ident \ TRM" -where + "CONS :: ident \ TRM" +as "Const" quotient_def - VAR :: "VAR :: name \ TRM" -where + "VAR :: name \ TRM" +as "Var" quotient_def - APP :: "APP :: TRM \ TRM \ TRM" -where + "APP :: TRM \ TRM \ TRM" +as "App" quotient_def - LAM :: "LAM :: TY \ name \ TRM \ TRM" -where + "LAM :: TY \ name \ TRM \ TRM" +as "Lam" thm TYP_def @@ -138,18 +138,18 @@ (* FIXME: print out a warning if the type contains a liftet type, like kind \ name set *) quotient_def - FV_kind :: "FV_kind :: KIND \ name set" -where + "FV_kind :: KIND \ name set" +as "fv_kind" quotient_def - FV_ty :: "FV_ty :: TY \ name set" -where + "FV_ty :: TY \ name set" +as "fv_ty" quotient_def - FV_trm :: "FV_trm :: TRM \ name set" -where + "FV_trm :: TRM \ name set" +as "fv_trm" thm FV_kind_def @@ -164,18 +164,18 @@ begin quotient_def - perm_kind :: "perm_kind :: 'x prm \ KIND \ KIND" -where + "perm_kind :: 'x prm \ KIND \ KIND" +as "(perm::'x prm \ kind \ kind)" quotient_def - perm_ty :: "perm_ty :: 'x prm \ TY \ TY" -where + "perm_ty :: 'x prm \ TY \ TY" +as "(perm::'x prm \ ty \ ty)" quotient_def - perm_trm :: "perm_trm :: 'x prm \ TRM \ TRM" -where + "perm_trm :: 'x prm \ TRM \ TRM" +as "(perm::'x prm \ trm \ trm)" end diff -r 6decb8811d30 -r 587e97d144a0 Quot/Examples/LamEx.thy --- a/Quot/Examples/LamEx.thy Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/Examples/LamEx.thy Fri Dec 11 11:14:05 2009 +0100 @@ -57,18 +57,18 @@ print_quotients quotient_def - Var :: "Var :: name \ lam" -where + "Var :: name \ lam" +as "rVar" quotient_def - App :: "App :: lam \ lam \ lam" -where + "App :: lam \ lam \ lam" +as "rApp" quotient_def - Lam :: "Lam :: name \ lam \ lam" -where + "Lam :: name \ lam \ lam" +as "rLam" thm Var_def @@ -76,8 +76,8 @@ thm Lam_def quotient_def - fv :: "fv :: lam \ name set" -where + "fv :: lam \ name set" +as "rfv" thm fv_def @@ -89,18 +89,12 @@ begin quotient_def - perm_lam :: "perm_lam :: 'x prm \ lam \ lam" -where + "perm_lam :: 'x prm \ lam \ lam" +as "perm::'x prm \ rlam \ rlam" end -(*quotient_def (for lam) - abs_fun_lam :: "'x prm \ lam \ lam" -where - "perm_lam \ (perm::'x prm \ rlam \ rlam)"*) - - thm perm_lam_def (* lemmas that need to lift *) diff -r 6decb8811d30 -r 587e97d144a0 Quot/Examples/LarryDatatype.thy --- a/Quot/Examples/LarryDatatype.thy Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/Examples/LarryDatatype.thy Fri Dec 11 11:14:05 2009 +0100 @@ -128,23 +128,23 @@ text{*The abstract message constructors*} quotient_def - Nonce::"Nonce :: nat \ msg" -where + "Nonce :: nat \ msg" +as "NONCE" quotient_def - MPair::"MPair :: msg \ msg \ msg" -where + "MPair :: msg \ msg \ msg" +as "MPAIR" quotient_def - Crypt::"Crypt :: nat \ msg \ msg" -where + "Crypt :: nat \ msg \ msg" +as "CRYPT" quotient_def - Decrypt::"Decrypt :: nat \ msg \ msg" -where + "Decrypt :: nat \ msg \ msg" +as "DECRYPT" lemma [quot_respect]: @@ -165,8 +165,8 @@ subsection{*The Abstract Function to Return the Set of Nonces*} quotient_def - nonces :: "nonces:: msg \ nat set" -where + "nonces:: msg \ nat set" +as "freenonces" text{*Now prove the four equations for @{term nonces}*} @@ -201,8 +201,8 @@ subsection{*The Abstract Function to Return the Left Part*} quotient_def - left :: "left:: msg \ msg" -where + "left:: msg \ msg" +as "freeleft" lemma [quot_respect]: @@ -228,8 +228,8 @@ subsection{*The Abstract Function to Return the Right Part*} quotient_def - right :: "right:: msg \ msg" -where + "right:: msg \ msg" +as "freeright" text{*Now prove the four equations for @{term right}*} @@ -362,8 +362,8 @@ need this function in order to prove discrimination theorems.*} quotient_def - discrim :: "discrim:: msg \ int" -where + "discrim:: msg \ int" +as "freediscrim" text{*Now prove the four equations for @{term discrim}*} diff -r 6decb8811d30 -r 587e97d144a0 Quot/quotient_def.ML --- a/Quot/quotient_def.ML Fri Dec 11 11:12:53 2009 +0100 +++ b/Quot/quotient_def.ML Fri Dec 11 11:14:05 2009 +0100 @@ -3,12 +3,9 @@ sig datatype flag = absF | repF val get_fun: flag -> Proof.context -> typ * typ -> term - val make_def: binding -> mixfix -> Attrib.binding -> term -> term -> + val make_def: mixfix -> Attrib.binding -> term -> term -> Proof.context -> (term * thm) * local_theory - - val quotdef: (binding * term * mixfix) * (Attrib.binding * term) -> - local_theory -> (term * thm) * local_theory - val quotdef_cmd: (binding * string * mixfix) * (Attrib.binding * string) -> + val quotdef_cmd: (Attrib.binding * string) * (mixfix * string) -> local_theory -> local_theory end; @@ -79,9 +76,11 @@ | (TVar _, TVar _) => raise (LIFT_MATCH "get_fun") | _ => raise (LIFT_MATCH "get_fun") -fun make_def qconst_bname mx attr lhs rhs lthy = +fun make_def mx attr lhs rhs lthy = let - val absrep_trm = get_fun absF lthy (fastype_of rhs, fastype_of lhs) $ rhs + val Free (lhs_str, lhs_ty) = lhs; + val qconst_bname = Binding.name lhs_str; + val absrep_trm = get_fun absF lthy (fastype_of rhs, lhs_ty) $ rhs |> Syntax.check_term lthy val prop = Logic.mk_equals (lhs, absrep_trm) val (_, prop') = LocalDefs.cert_def lthy prop @@ -91,12 +90,7 @@ fun qcinfo phi = qconsts_transfer phi {qconst = trm, rconst = rhs, def = thm} val lthy'' = Local_Theory.declaration true - (fn phi => - let - val qconst_str = Binding.name_of qconst_bname - in - qconsts_update_gen qconst_str (qcinfo phi) - end) lthy' + (fn phi => qconsts_update_gen lhs_str (qcinfo phi)) lthy' in ((trm, thm), lthy'') end @@ -111,22 +105,21 @@ (* - a meta-equation defining the constant, *) (* and the attributes of for this meta-equality *) -fun quotdef ((bind, lhs, mx), (attr, rhs)) lthy = - make_def bind mx attr lhs rhs lthy - -fun quotdef_cmd ((bind, lhsstr, mx), (attr, rhsstr)) lthy = +fun quotdef_cmd ((attr, lhsstr), (mx, rhsstr)) lthy = let val lhs = Syntax.read_term lthy lhsstr val rhs = Syntax.read_term lthy rhsstr in - quotdef ((bind, lhs, mx), (attr, rhs)) lthy |> snd + make_def mx attr lhs rhs lthy |> snd end +val _ = OuterKeyword.keyword "as"; + val quotdef_parser = - (OuterParse.binding -- - (OuterParse.$$$ "::" |-- OuterParse.!!! (OuterParse.term -- - OuterParse.opt_mixfix' --| OuterParse.where_)) >> OuterParse.triple2) -- - (SpecParse.opt_thm_name ":" -- OuterParse.term) + (SpecParse.opt_thm_name ":" -- + OuterParse.term) -- + (OuterParse.opt_mixfix' --| OuterParse.$$$ "as" -- + OuterParse.term) val _ = OuterSyntax.local_theory "quotient_def" "lifted definition of constants" OuterKeyword.thy_decl (quotdef_parser >> quotdef_cmd) diff -r 6decb8811d30 -r 587e97d144a0 isar-keywords-prove.el --- a/isar-keywords-prove.el Fri Dec 11 11:12:53 2009 +0100 +++ b/isar-keywords-prove.el Fri Dec 11 11:14:05 2009 +0100 @@ -257,6 +257,7 @@ (defconst isar-keywords-minor '("advanced" "and" + "as" "assumes" "attach" "avoids"