Nominal/Test.thy
changeset 1341 c25f797c7e6e
parent 1340 f201eb6acafc
child 1342 2b98012307f7
child 1353 3727e234fe6b
equal deleted inserted replaced
1340:f201eb6acafc 1341:c25f797c7e6e
     2 imports "Parser" "../Attic/Prove"
     2 imports "Parser" "../Attic/Prove"
     3 begin
     3 begin
     4 
     4 
     5 text {* weirdo example from Peter Sewell's bestiary *}
     5 text {* weirdo example from Peter Sewell's bestiary *}
     6 
     6 
     7 (*nominal_datatype weird =
     7 nominal_datatype weird =
     8   WBind x::"name" y::"name" p1::"weird" p2::"weird" p3::"weird"
     8   WBind x::"name" y::"name" p1::"weird" p2::"weird" p3::"weird"
     9     bind x in p1, bind x in p2, bind y in p2, bind y in p3
     9     bind x in p1, bind x in p2, bind y in p2, bind y in p3
    10 | WV "name"
    10 | WV "name"
    11 | WP "weird" "weird"
    11 | WP "weird" "weird"
    12 
    12 
    17 thm eqvts
    17 thm eqvts
    18 
    18 
    19 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding weird_inj}, []), (build_alpha_inj @{thms alpha_weird_raw.intros} @{thms weird_raw.distinct weird_raw.inject} @{thms alpha_weird_raw.cases} ctxt)) ctxt)) *}
    19 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding weird_inj}, []), (build_alpha_inj @{thms alpha_weird_raw.intros} @{thms weird_raw.distinct weird_raw.inject} @{thms alpha_weird_raw.cases} ctxt)) ctxt)) *}
    20 thm weird_inj
    20 thm weird_inj
    21 
    21 
    22 local_setup {*
    22 (*local_setup {*
    23 (fn ctxt => snd (Local_Theory.note ((@{binding alpha_eqvt}, []),
    23 (fn ctxt => snd (Local_Theory.note ((@{binding alpha_eqvt}, []),
    24 build_alpha_eqvts [@{term alpha_weird_raw}] [@{term "permute :: perm \<Rightarrow> weird_raw \<Rightarrow> weird_raw"}] @{thms permute_weird_raw.simps weird_inj} @{thm alpha_weird_raw.induct} ctxt) ctxt)) *}
    24 build_alpha_eqvts [@{term alpha_weird_raw}] [@{term "permute :: perm \<Rightarrow> weird_raw \<Rightarrow> weird_raw"}] @{thms permute_weird_raw.simps weird_inj} @{thm alpha_weird_raw.induct} ctxt) ctxt)) *}
    25 
    25 
    26 (*prove {* (snd o snd) (build_alpha_refl_gl [@{term alpha_weird_raw}] ("x","y","z")) *}
    26 (*prove {* (snd o snd) (build_alpha_refl_gl [@{term alpha_weird_raw}] ("x","y","z")) *}
    27 
    27 
   142 term Test.BP_raw
   142 term Test.BP_raw
   143 thm bi_raw.simps
   143 thm bi_raw.simps
   144 thm permute_lam_raw_permute_bp_raw.simps
   144 thm permute_lam_raw_permute_bp_raw.simps
   145 thm alpha_lam_raw_alpha_bp_raw.intros[no_vars]
   145 thm alpha_lam_raw_alpha_bp_raw.intros[no_vars]
   146 thm fv_lam_raw_fv_bp_raw.simps[no_vars]
   146 thm fv_lam_raw_fv_bp_raw.simps[no_vars]
   147 thm lam_bp_induct
   147 (*thm lam_bp_induct
   148 thm lam_bp_perm
   148 thm lam_bp_perm
   149 thm lam_bp_fv
   149 thm lam_bp_fv
   150 thm lam_bp_bn
   150 thm lam_bp_bn
       
   151 thm lam_bp_inject*)
   151 
   152 
   152 text {* example 2 *}
   153 text {* example 2 *}
   153 
   154 
   154 nominal_datatype trm' =
   155 nominal_datatype trm' =
   155   Var "name"
   156   Var "name"
   169 
   170 
   170 
   171 
   171 thm alpha_trm'_raw_alpha_pat'_raw.intros[no_vars]
   172 thm alpha_trm'_raw_alpha_pat'_raw.intros[no_vars]
   172 thm fv_trm'_raw_fv_pat'_raw.simps[no_vars]
   173 thm fv_trm'_raw_fv_pat'_raw.simps[no_vars]
   173 thm f_raw.simps
   174 thm f_raw.simps
   174 thm trm'_pat'_induct
   175 (*thm trm'_pat'_induct
   175 thm trm'_pat'_perm
   176 thm trm'_pat'_perm
   176 thm trm'_pat'_fv
   177 thm trm'_pat'_fv
   177 thm trm'_pat'_bn
   178 thm trm'_pat'_bn
       
   179 thm trm'_pat'_inject*)
   178 
   180 
   179 nominal_datatype trm0 =
   181 nominal_datatype trm0 =
   180   Var0 "name"
   182   Var0 "name"
   181 | App0 "trm0" "trm0"
   183 | App0 "trm0" "trm0"
   182 | Lam0 x::"name" t::"trm0"          bind x in t 
   184 | Lam0 x::"name" t::"trm0"          bind x in t 
   191   "f0 PN0 = {}"
   193   "f0 PN0 = {}"
   192 | "f0 (PS0 x) = {atom x}"
   194 | "f0 (PS0 x) = {atom x}"
   193 | "f0 (PD0 p1 p2) = (f0 p1) \<union> (f0 p2)"
   195 | "f0 (PD0 p1 p2) = (f0 p1) \<union> (f0 p2)"
   194 
   196 
   195 thm f0_raw.simps
   197 thm f0_raw.simps
   196 thm trm0_pat0_induct
   198 (*thm trm0_pat0_induct
   197 thm trm0_pat0_perm
   199 thm trm0_pat0_perm
   198 thm trm0_pat0_fv
   200 thm trm0_pat0_fv
   199 thm trm0_pat0_bn
   201 thm trm0_pat0_bn*)
   200 
   202 
   201 text {* example type schemes *}
   203 text {* example type schemes *}
   202 
   204 
   203 (* does not work yet
   205 (* does not work yet
   204 nominal_datatype t =
   206 nominal_datatype t =
   208 nominal_datatype tyS =
   210 nominal_datatype tyS =
   209   All xs::"name list" ty::"t_raw" bind xs in ty
   211   All xs::"name list" ty::"t_raw" bind xs in ty
   210 *)
   212 *)
   211 
   213 
   212 
   214 
   213 (* alpha_eqvt fails...
       
   214 nominal_datatype t = 
   215 nominal_datatype t = 
   215   Var "name" 
   216   Var "name" 
   216 | Fun "t" "t"
   217 | Fun "t" "t"
   217 and  tyS = 
   218 and  tyS = 
   218   All xs::"name set" ty::"t" bind xs in ty *)
   219   All xs::"name set" ty::"t" bind xs in ty
   219 
   220 
   220 (* example 1 from Terms.thy *)
   221 (* example 1 from Terms.thy *)
   221 
   222 
   222 nominal_datatype trm1 =
   223 nominal_datatype trm1 =
   223   Vr1 "name"
   224   Vr1 "name"
   269 
   270 
   270 (* example 4 from Terms.thy *)
   271 (* example 4 from Terms.thy *)
   271 
   272 
   272 (* fv_eqvt does not work, we need to repaire defined permute functions
   273 (* fv_eqvt does not work, we need to repaire defined permute functions
   273    defined fv and defined alpha... *)
   274    defined fv and defined alpha... *)
   274 (*nominal_datatype trm4 =
   275 nominal_datatype trm4 =
   275   Vr4 "name"
   276   Vr4 "name"
   276 | Ap4 "trm4" "trm4 list"
   277 | Ap4 "trm4" "trm4 list"
   277 | Lm4 x::"name" t::"trm4"  bind x in t
   278 | Lm4 x::"name" t::"trm4"  bind x in t
   278 
   279 
   279 thm alpha_trm4_raw_alpha_trm4_raw_list.intros[no_vars]
   280 thm alpha_trm4_raw_alpha_trm4_raw_list.intros[no_vars]
   280 thm fv_trm4_raw_fv_trm4_raw_list.simps[no_vars]*)
   281 thm fv_trm4_raw_fv_trm4_raw_list.simps[no_vars]
   281 
   282 
   282 (* example 5 from Terms.thy *)
   283 (* example 5 from Terms.thy *)
   283 
   284 
   284 nominal_datatype trm5 =
   285 nominal_datatype trm5 =
   285   Vr5 "name"
   286   Vr5 "name"
   294   "bv5 Lnil = {}"
   295   "bv5 Lnil = {}"
   295 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)"
   296 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)"
   296 
   297 
   297 (* example 6 from Terms.thy *)
   298 (* example 6 from Terms.thy *)
   298 
   299 
   299 (* BV is not respectful, needs to fail
   300 (* BV is not respectful, needs to fail*)
   300 nominal_datatype trm6 =
   301 nominal_datatype trm6 =
   301   Vr6 "name"
   302   Vr6 "name"
   302 | Lm6 x::"name" t::"trm6"         bind x in t
   303 | Lm6 x::"name" t::"trm6"         bind x in t
   303 | Lt6 left::"trm6" right::"trm6"  bind "bv6 left" in right
   304 | Lt6 left::"trm6" right::"trm6"  bind "bv6 left" in right
   304 binder
   305 binder
   305   bv6
   306   bv6
   306 where
   307 where
   307   "bv6 (Vr6 n) = {}"
   308   "bv6 (Vr6 n) = {}"
   308 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t"
   309 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t"
   309 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" *)
   310 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r"
   310 
       
   311 (* example 7 from Terms.thy *)
   311 (* example 7 from Terms.thy *)
   312 
   312 
   313 (* BV is not respectful, needs to fail
   313 (* BV is not respectful, needs to fail*)
   314 nominal_datatype trm7 =
   314 nominal_datatype trm7 =
   315   Vr7 "name"
   315   Vr7 "name"
   316 | Lm7 l::"name" r::"trm7"   bind l in r
   316 | Lm7 l::"name" r::"trm7"   bind l in r
   317 | Lt7 l::"trm7" r::"trm7"   bind "bv7 l" in r
   317 | Lt7 l::"trm7" r::"trm7"   bind "bv7 l" in r
   318 binder 
   318 binder 
   319   bv7 
   319   bv7 
   320 where
   320 where
   321   "bv7 (Vr7 n) = {atom n}"
   321   "bv7 (Vr7 n) = {atom n}"
   322 | "bv7 (Lm7 n t) = bv7 t - {atom n}"
   322 | "bv7 (Lm7 n t) = bv7 t - {atom n}"
   323 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" *)
   323 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r"
   324 
   324 
   325 (* example 8 from Terms.thy *)
   325 (* example 8 from Terms.thy *)
   326 
   326 
   327 nominal_datatype foo8 =
   327 nominal_datatype foo8 =
   328   Foo0 "name"
   328   Foo0 "name"
   336   "bv8 (Bar0 x) = {}"
   336   "bv8 (Bar0 x) = {}"
   337 | "bv8 (Bar1 v x b) = {atom v}"
   337 | "bv8 (Bar1 v x b) = {atom v}"
   338 
   338 
   339 (* example 9 from Terms.thy *)
   339 (* example 9 from Terms.thy *)
   340 
   340 
   341 (* BV is not respectful, needs to fail
   341 (* BV is not respectful, needs to fail*)
   342 nominal_datatype lam9 =
   342 nominal_datatype lam9 =
   343   Var9 "name"
   343   Var9 "name"
   344 | Lam9 n::"name" l::"lam9" bind n in l
   344 | Lam9 n::"name" l::"lam9" bind n in l
   345 and bla9 =
   345 and bla9 =
   346   Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s
   346   Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s
   347 binder
   347 binder
   348   bv9
   348   bv9
   349 where
   349 where
   350   "bv9 (Var9 x) = {}"
   350   "bv9 (Var9 x) = {}"
   351 | "bv9 (Lam9 x b) = {atom x}" *)
   351 | "bv9 (Lam9 x b) = {atom x}"
   352 
   352 
   353 (* example from my PHD *)
   353 (* example from my PHD *)
   354 
   354 
   355 atom_decl coname
   355 atom_decl coname
   356 
   356 
   357 (*nominal_datatype phd =
   357 nominal_datatype phd =
   358    Ax "name" "coname"
   358    Ax "name" "coname"
   359 |  Cut n::"coname" t1::"phd" c::"coname" t2::"phd"              bind n in t1, bind c in t2
   359 |  Cut n::"coname" t1::"phd" c::"coname" t2::"phd"              bind n in t1, bind c in t2
   360 |  AndR c1::"coname" t1::"phd" c2::"coname" t2::"phd" "coname"  bind c1 in t1, bind c2 in t2
   360 |  AndR c1::"coname" t1::"phd" c2::"coname" t2::"phd" "coname"  bind c1 in t1, bind c2 in t2
   361 |  AndL1 n::"name" t::"phd" "name"                              bind n in t
   361 |  AndL1 n::"name" t::"phd" "name"                              bind n in t
   362 |  AndL2 n::"name" t::"phd" "name"                              bind n in t
   362 |  AndL2 n::"name" t::"phd" "name"                              bind n in t
   363 |  ImpL c::"coname" t1::"phd" n::"name" t2::"phd" "name"        bind c in t1, bind n in t2
   363 |  ImpL c::"coname" t1::"phd" n::"name" t2::"phd" "name"        bind c in t1, bind n in t2
   364 |  ImpR c::"coname" n::"name" t::"phd" "coname"                 bind n in t, bind c in t
   364 |  ImpR c::"coname" n::"name" t::"phd" "coname"                 bind n in t, bind c in t
   365 
   365 
   366 thm alpha_phd_raw.intros[no_vars]
   366 thm alpha_phd_raw.intros[no_vars]
   367 thm fv_phd_raw.simps[no_vars]
   367 thm fv_phd_raw.simps[no_vars]
   368 *)
   368 
   369 
   369 
   370 (* example form Leroy 96 about modules; OTT *)
   370 (* example form Leroy 96 about modules; OTT *)
   371 
   371 
   372 nominal_datatype mexp =
   372 nominal_datatype mexp =
   373   Acc "path"
   373   Acc "path"
   496 | "bp (PUnit) = {}"
   496 | "bp (PUnit) = {}"
   497 | "bp (PPair p1 p2) = bp p1 \<union> bp p2"
   497 | "bp (PPair p1 p2) = bp p1 \<union> bp p2"
   498 
   498 
   499 thm quot_respect
   499 thm quot_respect
   500 (* example 6 from Peter Sewell's bestiary *)
   500 (* example 6 from Peter Sewell's bestiary *)
   501 (*nominal_datatype exp6 =
   501 nominal_datatype exp6 =
   502   EVar name
   502   EVar name
   503 | EPair exp6 exp6
   503 | EPair exp6 exp6
   504 | ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1
   504 | ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1
   505 and pat6 =
   505 and pat6 =
   506   PVar' name
   506   PVar' name
   509 binder
   509 binder
   510   bp6 :: "pat6 \<Rightarrow> atom set"
   510   bp6 :: "pat6 \<Rightarrow> atom set"
   511 where
   511 where
   512   "bp6 (PVar' x) = {atom x}"
   512   "bp6 (PVar' x) = {atom x}"
   513 | "bp6 (PUnit') = {}"
   513 | "bp6 (PUnit') = {}"
   514 | "bp6 (PPair' p1 p2) = bp6 p1 \<union> bp6 p2"*)
   514 | "bp6 (PPair' p1 p2) = bp6 p1 \<union> bp6 p2"
   515 
   515 
   516 (* example 7 from Peter Sewell's bestiary *)
   516 (* example 7 from Peter Sewell's bestiary *)
   517 (*nominal_datatype exp7 =
   517 nominal_datatype exp7 =
   518   EVar name
   518   EVar name
   519 | EUnit
   519 | EUnit
   520 | EPair exp7 exp7
   520 | EPair exp7 exp7
   521 | ELetRec l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l
   521 | ELetRec l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l
   522 and lrb =
   522 and lrb =
   528   b7 :: "lrb \<Rightarrow> atom set" and
   528   b7 :: "lrb \<Rightarrow> atom set" and
   529   b7s :: "lrbs \<Rightarrow> atom set"
   529   b7s :: "lrbs \<Rightarrow> atom set"
   530 where
   530 where
   531   "b7 (Assign x e) = {atom x}"
   531   "b7 (Assign x e) = {atom x}"
   532 | "b7s (Single a) = b7 a"
   532 | "b7s (Single a) = b7 a"
   533 | "b7s (More a as) = (b7 a) \<union> (b7s as)"*)
   533 | "b7s (More a as) = (b7 a) \<union> (b7s as)"
   534 
   534 
   535 (* example 8 from Peter Sewell's bestiary *)
   535 (* example 8 from Peter Sewell's bestiary *)
   536 (*nominal_datatype exp8 =
   536 nominal_datatype exp8 =
   537   EVar name
   537   EVar name
   538 | EUnit
   538 | EUnit
   539 | EPair exp8 exp8
   539 | EPair exp8 exp8
   540 | ELetRec l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l
   540 | ELetRec l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l
   541 and fnclause =
   541 and fnclause =
   565 | "b_pat (PUnit) = {}"
   565 | "b_pat (PUnit) = {}"
   566 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
   566 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
   567 | "b_fnclauses (S fc) = (b_fnclause fc)"
   567 | "b_fnclauses (S fc) = (b_fnclause fc)"
   568 | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)"
   568 | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)"
   569 | "b_lrb8 (Clause fcs) = (b_fnclauses fcs)"
   569 | "b_lrb8 (Clause fcs) = (b_fnclauses fcs)"
   570 | "b_fnclause (K x pat exp8) = {atom x}"*)
   570 | "b_fnclause (K x pat exp8) = {atom x}"
   571 
   571 
   572 
   572 
   573 
   573 
   574 
   574 
   575 (* example 9 from Peter Sewell's bestiary *)
   575 (* example 9 from Peter Sewell's bestiary *)