# HG changeset patch # User Cezary Kaliszyk # Date 1265987050 -3600 # Node ID c4001cda9da3ec1f436695a64aba4b66c3a63b7b # Parent 93c9022ba5e9f34bfb43e3138ab9f007ac21e8d4 renamed 'as' to 'is' everywhere. diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/FSet.thy --- a/Quot/Examples/FSet.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/FSet.thy Fri Feb 12 16:04:10 2010 +0100 @@ -28,17 +28,17 @@ quotient_definition "EMPTY :: 'a fset" -as +is "[]::'a list" quotient_definition "INSERT :: 'a \ 'a fset \ 'a fset" -as +is "op #" quotient_definition "FUNION :: 'a fset \ 'a fset \ 'a fset" -as +is "op @" fun @@ -49,12 +49,12 @@ quotient_definition "CARD :: 'a fset \ nat" -as +is "card1" quotient_definition "fconcat :: ('a fset) fset \ 'a fset" -as +is "concat" term concat @@ -120,17 +120,17 @@ quotient_definition "IN :: 'a \ 'a fset \ bool" -as +is "op mem" quotient_definition "FOLD :: ('a \ 'a \ 'a) \ ('b \ 'a) \ 'a \ 'b fset \ 'a" -as +is "fold1" quotient_definition "fmap :: ('a \ 'b) \ 'a fset \ 'b fset" -as +is "map" lemma mem_rsp: @@ -285,12 +285,12 @@ quotient_definition "EMPTY2 :: 'a fset2" -as +is "[]::'a list" quotient_definition "INSERT2 :: 'a \ 'a fset2 \ 'a fset2" -as +is "op #" lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" @@ -303,12 +303,12 @@ quotient_definition "fset_rec :: 'a \ ('b \ 'b fset \ 'a \ 'a) \ 'b fset \ 'a" -as +is "list_rec" quotient_definition "fset_case :: 'a \ ('b \ 'b fset \ 'a) \ 'b fset \ 'a" -as +is "list_case" (* Probably not true without additional assumptions about the function *) diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/FSet2.thy --- a/Quot/Examples/FSet2.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/FSet2.thy Fri Feb 12 16:04:10 2010 +0100 @@ -27,12 +27,12 @@ quotient_definition "fempty :: 'a fset" ("{||}") -as +is "[]" quotient_definition "finsert :: 'a \ 'a fset \ 'a fset" -as +is "(op #)" lemma finsert_rsp[quot_respect]: @@ -41,7 +41,7 @@ quotient_definition "funion :: 'a fset \ 'a fset \ 'a fset" ("_ \f _" [65,66] 65) -as +is "(op @)" lemma append_rsp_aux1: @@ -68,7 +68,7 @@ quotient_definition "fmem :: 'a \ 'a fset \ bool" ("_ \f _" [50, 51] 50) -as +is "(op mem)" lemma memb_rsp_aux: @@ -92,7 +92,7 @@ quotient_definition "finter::'a fset \ 'a fset \ 'a fset" ("_ \f _" [70, 71] 70) -as +is "inter_list" no_syntax diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/FSet3.thy --- a/Quot/Examples/FSet3.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/FSet3.thy Fri Feb 12 16:04:10 2010 +0100 @@ -136,7 +136,7 @@ quotient_definition "funion :: 'a fset \ 'a fset \ 'a fset" (infixl "|\|" 65) -as +is "op @" section {* Cardinality of finite sets *} @@ -227,12 +227,12 @@ quotient_definition "fmap :: ('a \ 'b) \ 'a fset \ 'b fset" -as +is "map" quotient_definition "fconcat :: ('a fset) fset => 'a fset" -as +is "concat" (*lemma fconcat_rsp[quot_respect]: @@ -619,19 +619,19 @@ quotient_definition "fdelete :: 'a fset \ 'a \ 'a fset" - as "delete_raw" + is "delete_raw" quotient_definition "finter :: 'a fset \ 'a fset \ 'a fset" ("_ \f _" [70, 71] 70) - as "inter_raw" + is "inter_raw" quotient_definition "ffold :: ('a \ 'b \ 'b) \ 'b \ 'a fset \ 'b" - as "fold_raw" + is "fold_raw" quotient_definition "fset_to_set :: 'a fset \ 'a set" - as "set" + is "set" section {* Lifted Theorems *} @@ -690,7 +690,7 @@ quotient_definition "fset_case :: 'a \ ('b \ 'b fset \ 'a) \ 'b fset \ 'a" -as +is "list_case" (* NOT SURE IF TRUE *) diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/IntEx.thy --- a/Quot/Examples/IntEx.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/IntEx.thy Fri Feb 12 16:04:10 2010 +0100 @@ -15,12 +15,12 @@ quotient_definition "ZERO :: my_int" -as +is "(0::nat, 0::nat)" quotient_definition "ONE :: my_int" -as +is "(1::nat, 0::nat)" fun @@ -30,7 +30,7 @@ quotient_definition "PLUS :: my_int \ my_int \ my_int" -as +is "my_plus" fun @@ -40,7 +40,7 @@ quotient_definition "NEG :: my_int \ my_int" -as +is "my_neg" definition @@ -55,7 +55,7 @@ quotient_definition "MULT :: my_int \ my_int \ my_int" -as +is "my_mult" @@ -67,7 +67,7 @@ quotient_definition "LE :: my_int \ my_int \ bool" -as +is "my_le" term LE diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/IntEx2.thy --- a/Quot/Examples/IntEx2.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/IntEx2.thy Fri Feb 12 16:04:10 2010 +0100 @@ -21,10 +21,10 @@ ML {* @{term "0 \ int"} *} quotient_definition - "0 \ int" as "(0\nat, 0\nat)" + "0 \ int" is "(0\nat, 0\nat)" quotient_definition - "1 \ int" as "(1\nat, 0\nat)" + "1 \ int" is "(1\nat, 0\nat)" fun plus_raw :: "(nat \ nat) \ (nat \ nat) \ (nat \ nat)" @@ -32,7 +32,7 @@ "plus_raw (x, y) (u, v) = (x + u, y + v)" quotient_definition - "(op +) \ (int \ int \ int)" as "plus_raw" + "(op +) \ (int \ int \ int)" is "plus_raw" fun uminus_raw :: "(nat \ nat) \ (nat \ nat)" @@ -40,7 +40,7 @@ "uminus_raw (x, y) = (y, x)" quotient_definition - "(uminus \ (int \ int))" as "uminus_raw" + "(uminus \ (int \ int))" is "uminus_raw" definition minus_int_def [code del]: "z - w = z + (-w\int)" @@ -51,7 +51,7 @@ "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)" quotient_definition - mult_int_def: "(op *) :: (int \ int \ int)" as "mult_raw" + mult_int_def: "(op *) :: (int \ int \ int)" is "mult_raw" fun le_raw :: "(nat \ nat) \ (nat \ nat) \ bool" @@ -59,7 +59,7 @@ "le_raw (x, y) (u, v) = (x+v \ u+y)" quotient_definition - le_int_def: "(op \) :: int \ int \ bool" as "le_raw" + le_int_def: "(op \) :: int \ int \ bool" is "le_raw" definition less_int_def [code del]: "(z\int) < w = (z \ w \ z \ w)" @@ -207,7 +207,7 @@ "int_of_nat_raw m = (m :: nat, 0 :: nat)" quotient_definition - "int_of_nat :: nat \ int" as "int_of_nat_raw" + "int_of_nat :: nat \ int" is "int_of_nat_raw" lemma[quot_respect]: shows "(op = ===> op \) int_of_nat_raw int_of_nat_raw" diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/LFex.thy --- a/Quot/Examples/LFex.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/LFex.thy Fri Feb 12 16:04:10 2010 +0100 @@ -83,48 +83,48 @@ quotient_definition "TYP :: KIND" -as +is "Type" quotient_definition "KPI :: TY \ name \ KIND \ KIND" -as +is "KPi" quotient_definition "TCONST :: ident \ TY" -as +is "TConst" quotient_definition "TAPP :: TY \ TRM \ TY" -as +is "TApp" quotient_definition "TPI :: TY \ name \ TY \ TY" -as +is "TPi" (* FIXME: does not work with CONST *) quotient_definition "CONS :: ident \ TRM" -as +is "Const" quotient_definition "VAR :: name \ TRM" -as +is "Var" quotient_definition "APP :: TRM \ TRM \ TRM" -as +is "App" quotient_definition "LAM :: TY \ name \ TRM \ TRM" -as +is "Lam" thm TYP_def @@ -140,17 +140,17 @@ (* FIXME: print out a warning if the type contains a liftet type, like kind \ name set *) quotient_definition "FV_kind :: KIND \ name set" -as +is "fv_kind" quotient_definition "FV_ty :: TY \ name set" -as +is "fv_ty" quotient_definition "FV_trm :: TRM \ name set" -as +is "fv_trm" thm FV_kind_def @@ -166,17 +166,17 @@ quotient_definition "perm_kind :: 'x prm \ KIND \ KIND" -as +is "(perm::'x prm \ kind \ kind)" quotient_definition "perm_ty :: 'x prm \ TY \ TY" -as +is "(perm::'x prm \ ty \ ty)" quotient_definition "perm_trm :: 'x prm \ TRM \ TRM" -as +is "(perm::'x prm \ trm \ trm)" end diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/LamEx.thy --- a/Quot/Examples/LamEx.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/LamEx.thy Fri Feb 12 16:04:10 2010 +0100 @@ -178,22 +178,22 @@ quotient_definition "Var :: name \ lam" -as +is "rVar" quotient_definition "App :: lam \ lam \ lam" -as +is "rApp" quotient_definition "Lam :: name \ lam \ lam" -as +is "rLam" quotient_definition "fv :: lam \ name set" -as +is "rfv" (* definition of overloaded permutation function *) @@ -204,7 +204,7 @@ quotient_definition "perm_lam :: 'x prm \ lam \ lam" -as +is "perm::'x prm \ rlam \ rlam" end diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/LarryDatatype.thy --- a/Quot/Examples/LarryDatatype.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/LarryDatatype.thy Fri Feb 12 16:04:10 2010 +0100 @@ -128,22 +128,22 @@ quotient_definition "Nonce :: nat \ msg" -as +is "NONCE" quotient_definition "MPair :: msg \ msg \ msg" -as +is "MPAIR" quotient_definition "Crypt :: nat \ msg \ msg" -as +is "CRYPT" quotient_definition "Decrypt :: nat \ msg \ msg" -as +is "DECRYPT" lemma [quot_respect]: @@ -167,7 +167,7 @@ quotient_definition "nonces:: msg \ nat set" -as +is "freenonces" text{*Now prove the four equations for @{term nonces}*} @@ -204,7 +204,7 @@ quotient_definition "left:: msg \ msg" -as +is "freeleft" lemma [quot_respect]: @@ -231,7 +231,7 @@ quotient_definition "right:: msg \ msg" -as +is "freeright" text{*Now prove the four equations for @{term right}*} @@ -365,7 +365,7 @@ quotient_definition "discrim:: msg \ int" -as +is "freediscrim" text{*Now prove the four equations for @{term discrim}*} diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/LarryInt.thy --- a/Quot/Examples/LarryInt.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/LarryInt.thy Fri Feb 12 16:04:10 2010 +0100 @@ -16,18 +16,18 @@ instantiation int :: "{zero, one, plus, uminus, minus, times, ord}" begin -quotient_definition - Zero_int_def: "0::int" as "(0::nat, 0::nat)" +quotient_definition + Zero_int_def: "0::int" is "(0::nat, 0::nat)" quotient_definition - One_int_def: "1::int" as "(1::nat, 0::nat)" + One_int_def: "1::int" is "(1::nat, 0::nat)" definition "add_raw \ \(x, y) (u, v). (x + (u::nat), y + (v::nat))" quotient_definition "(op +) :: int \ int \ int" -as +is "add_raw" definition @@ -35,7 +35,7 @@ quotient_definition "uminus :: int \ int" -as +is "uminus_raw" fun @@ -45,7 +45,7 @@ quotient_definition "(op *) :: int \ int \ int" -as +is "mult_raw" definition @@ -53,7 +53,7 @@ quotient_definition le_int_def: "(op \) :: int \ int \ bool" -as +is "le_raw" definition @@ -369,7 +369,7 @@ quotient_definition "nat2::int \ nat" -as +is "nat_raw" abbreviation diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Examples/SigmaEx.thy --- a/Quot/Examples/SigmaEx.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Examples/SigmaEx.thy Fri Feb 12 16:04:10 2010 +0100 @@ -35,27 +35,27 @@ quotient_definition "Var :: name \ obj" -as +is "rVar" quotient_definition "Obj :: (string \ method) list \ obj" -as +is "rObj" quotient_definition "Inv :: obj \ string \ obj" -as +is "rInv" quotient_definition "Upd :: obj \ string \ method \ obj" -as +is "rUpd" quotient_definition "Sig :: name \ obj \ method" -as +is "rSig" overloading @@ -65,12 +65,12 @@ quotient_definition "perm_obj :: 'x prm \ obj \ obj" -as +is "(perm::'x prm \ robj \ robj)" quotient_definition "perm_method :: 'x prm \ method \ method" -as +is "(perm::'x prm \ rmethod \ rmethod)" end diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Nominal/Abs.thy --- a/Quot/Nominal/Abs.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Nominal/Abs.thy Fri Feb 12 16:04:10 2010 +0100 @@ -194,9 +194,9 @@ done quotient_definition - "Abs::atom set \ ('a::pt) \ 'a abs" -as - "Pair::atom set \ ('a::pt) \ (atom set \ 'a)" + "Abs::atom set \ ('a::pt) \ 'a abs" +is + "Pair::atom set \ ('a::pt) \ (atom set \ 'a)" lemma [quot_respect]: shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair" @@ -234,7 +234,7 @@ quotient_definition "permute_abs::perm \ ('a::pt abs) \ 'a abs" -as +is "permute:: perm \ (atom set \ 'a::pt) \ (atom set \ 'a::pt)" lemma permute_ABS [simp]: @@ -252,7 +252,7 @@ quotient_definition "supp_Abs_fun :: ('a::pt) abs \ atom \ bool" -as +is "supp_abs_fun" lemma supp_Abs_fun_simp: diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Nominal/LFex.thy --- a/Quot/Nominal/LFex.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Nominal/LFex.thy Fri Feb 12 16:04:10 2010 +0100 @@ -256,64 +256,64 @@ quotient_definition "TYP :: KIND" -as +is "Type" quotient_definition "KPI :: TY \ name \ KIND \ KIND" -as +is "KPi" quotient_definition "TCONST :: ident \ TY" -as +is "TConst" quotient_definition "TAPP :: TY \ TRM \ TY" -as +is "TApp" quotient_definition "TPI :: TY \ name \ TY \ TY" -as +is "TPi" (* FIXME: does not work with CONST *) quotient_definition "CONS :: ident \ TRM" -as +is "Const" quotient_definition "VAR :: name \ TRM" -as +is "Var" quotient_definition "APP :: TRM \ TRM \ TRM" -as +is "App" quotient_definition "LAM :: TY \ name \ TRM \ TRM" -as +is "Lam" (* FIXME: print out a warning if the type contains a liftet type, like kind \ name set *) quotient_definition "fv_kind :: KIND \ atom set" -as +is "rfv_kind" quotient_definition "fv_ty :: TY \ atom set" -as +is "rfv_ty" quotient_definition "fv_trm :: TRM \ atom set" -as +is "rfv_trm" lemma alpha_rfv: @@ -390,17 +390,17 @@ quotient_definition "permute_KIND :: perm \ KIND \ KIND" -as +is "permute :: perm \ kind \ kind" quotient_definition "permute_TY :: perm \ TY \ TY" -as +is "permute :: perm \ ty \ ty" quotient_definition "permute_TRM :: perm \ TRM \ TRM" -as +is "permute :: perm \ trm \ trm" lemmas permute_ktt[simp] = permute_kind_permute_ty_permute_trm.simps[quot_lifted] diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Nominal/LamEx.thy --- a/Quot/Nominal/LamEx.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Nominal/LamEx.thy Fri Feb 12 16:04:10 2010 +0100 @@ -282,22 +282,22 @@ quotient_definition "Var :: name \ lam" -as +is "rVar" quotient_definition "App :: lam \ lam \ lam" -as +is "rApp" quotient_definition "Lam :: name \ lam \ lam" -as +is "rLam" quotient_definition "fv :: lam \ atom set" -as +is "rfv" lemma perm_rsp[quot_respect]: @@ -346,7 +346,7 @@ quotient_definition "permute_lam :: perm \ lam \ lam" -as +is "permute :: perm \ rlam \ rlam" lemma permute_lam [simp]: diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Nominal/LamEx2.thy --- a/Quot/Nominal/LamEx2.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Nominal/LamEx2.thy Fri Feb 12 16:04:10 2010 +0100 @@ -175,22 +175,22 @@ quotient_definition "Var :: name \ lam" -as +is "rVar" quotient_definition "App :: lam \ lam \ lam" -as +is "rApp" quotient_definition "Lam :: name \ lam \ lam" -as +is "rLam" quotient_definition "fv :: lam \ atom set" -as +is "rfv" lemma perm_rsp[quot_respect]: @@ -239,7 +239,7 @@ quotient_definition "permute_lam :: perm \ lam \ lam" -as +is "permute :: perm \ rlam \ rlam" lemma permute_lam [simp]: diff -r 93c9022ba5e9 -r c4001cda9da3 Quot/Nominal/Terms.thy --- a/Quot/Nominal/Terms.thy Fri Feb 12 15:50:43 2010 +0100 +++ b/Quot/Nominal/Terms.thy Fri Feb 12 16:04:10 2010 +0100 @@ -150,27 +150,27 @@ quotient_definition "Vr1 :: name \ trm1" -as +is "rVr1" quotient_definition "Ap1 :: trm1 \ trm1 \ trm1" -as +is "rAp1" quotient_definition "Lm1 :: name \ trm1 \ trm1" -as +is "rLm1" quotient_definition "Lt1 :: bp \ trm1 \ trm1 \ trm1" -as +is "rLt1" quotient_definition "fv_trm1 :: trm1 \ atom set" -as +is "rfv_trm1" lemma alpha_rfv1: @@ -211,7 +211,7 @@ quotient_definition "permute_trm1 :: perm \ trm1 \ trm1" -as +is "permute :: perm \ rtrm1 \ rtrm1" lemmas permute_trm1[simp] = permute_rtrm1_permute_bp.simps[quot_lifted] @@ -705,42 +705,42 @@ quotient_definition "Vr5 :: name \ trm5" -as +is "rVr5" quotient_definition "Ap5 :: trm5 \ trm5 \ trm5" -as +is "rAp5" quotient_definition "Lt5 :: lts \ trm5 \ trm5" -as +is "rLt5" quotient_definition "Lnil :: lts" -as +is "rLnil" quotient_definition "Lcons :: name \ trm5 \ lts \ lts" -as +is "rLcons" quotient_definition "fv_trm5 :: trm5 \ atom set" -as +is "rfv_trm5" quotient_definition "fv_lts :: lts \ atom set" -as +is "rfv_lts" quotient_definition "bv5 :: lts \ atom set" -as +is "rbv5" lemma rbv5_eqvt: @@ -827,12 +827,12 @@ quotient_definition "permute_trm5 :: perm \ trm5 \ trm5" -as +is "permute :: perm \ rtrm5 \ rtrm5" quotient_definition "permute_lts :: perm \ lts \ lts" -as +is "permute :: perm \ rlts \ rlts" lemma trm5_lts_zero: @@ -995,27 +995,27 @@ quotient_definition "Vr6 :: name \ trm6" -as +is "rVr6" quotient_definition "Lm6 :: name \ trm6 \ trm6" -as +is "rLm6" quotient_definition "Lt6 :: trm6 \ trm6 \ trm6" -as +is "rLt6" quotient_definition "fv_trm6 :: trm6 \ atom set" -as +is "rfv_trm6" quotient_definition "bv6 :: trm6 \ atom set" -as +is "rbv6" lemma [quot_respect]: @@ -1077,7 +1077,7 @@ quotient_definition "permute_trm6 :: perm \ trm6 \ trm6" -as +is "permute :: perm \ rtrm6 \ rtrm6" instance @@ -1316,32 +1316,32 @@ quotient_definition "qVar9 :: name \ lam9" -as +is "Var9" quotient_definition "qLam :: name \ lam9 \ lam9" -as +is "Lam9" quotient_definition "qBla9 :: lam9 \ lam9 \ bla9" -as +is "Bla9" quotient_definition "fv_lam9 :: lam9 \ atom set" -as +is "rfv_lam9" quotient_definition "fv_bla9 :: bla9 \ atom set" -as +is "rfv_bla9" quotient_definition "bv9 :: lam9 \ atom set" -as +is "rbv9" instantiation lam9 and bla9 :: pt @@ -1349,12 +1349,12 @@ quotient_definition "permute_lam9 :: perm \ lam9 \ lam9" -as +is "permute :: perm \ rlam9 \ rlam9" quotient_definition "permute_bla9 :: perm \ bla9 \ bla9" -as +is "permute :: perm \ rbla9 \ rbla9" instance