Nominal/Parser.thy
changeset 2382 e8b9c0ebf5dd
parent 2381 fd85f4921654
child 2383 83f1b16486ee
equal deleted inserted replaced
2381:fd85f4921654 2382:e8b9c0ebf5dd
     1 theory NewParser
       
     2 imports "../Nominal-General/Nominal2_Base" 
       
     3         "../Nominal-General/Nominal2_Eqvt" 
       
     4         "../Nominal-General/Nominal2_Supp" 
       
     5         "Perm" "NewFv" "NewAlpha" "Tacs" "Equivp" "Lift"
       
     6 begin
       
     7 
       
     8 section{* Interface for nominal_datatype *}
       
     9 
       
    10 
       
    11 ML {* 
       
    12 (* nominal datatype parser *)
       
    13 local
       
    14   structure P = OuterParse;
       
    15   structure S = Scan
       
    16 
       
    17   fun triple1 ((x, y), z) = (x, y, z)
       
    18   fun triple2 (x, (y, z)) = (x, y, z)
       
    19   fun tuple ((x, y, z), u) = (x, y, z, u)
       
    20   fun tswap (((x, y), z), u) = (x, y, u, z)
       
    21 in
       
    22 
       
    23 val _ = OuterKeyword.keyword "bind"
       
    24 val _ = OuterKeyword.keyword "bind_set"
       
    25 val _ = OuterKeyword.keyword "bind_res"
       
    26 
       
    27 val anno_typ = S.option (P.name --| P.$$$ "::") -- P.typ
       
    28 
       
    29 val bind_mode = P.$$$ "bind" || P.$$$ "bind_set" || P.$$$ "bind_res"
       
    30 
       
    31 val bind_clauses = 
       
    32   P.enum "," (bind_mode -- S.repeat1 P.term -- (P.$$$ "in" |-- S.repeat1 P.name) >> triple1)
       
    33 
       
    34 val cnstr_parser =
       
    35   P.binding -- S.repeat anno_typ -- bind_clauses -- P.opt_mixfix >> tswap
       
    36 
       
    37 (* datatype parser *)
       
    38 val dt_parser =
       
    39   (P.type_args -- P.binding -- P.opt_mixfix >> triple1) -- 
       
    40     (P.$$$ "=" |-- P.enum1 "|" cnstr_parser) >> tuple
       
    41 
       
    42 (* binding function parser *)
       
    43 val bnfun_parser = 
       
    44   S.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([], [])
       
    45 
       
    46 (* main parser *)
       
    47 val main_parser =
       
    48   P.and_list1 dt_parser -- bnfun_parser >> triple2
       
    49 
       
    50 end
       
    51 *}
       
    52 
       
    53 ML {*
       
    54 fun get_cnstrs dts =
       
    55   map (fn (_, _, _, constrs) => constrs) dts
       
    56 
       
    57 fun get_typed_cnstrs dts =
       
    58   flat (map (fn (_, bn, _, constrs) => 
       
    59    (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts)
       
    60 
       
    61 fun get_cnstr_strs dts =
       
    62   map (fn (bn, _, _) => Binding.name_of bn) (flat (get_cnstrs dts))
       
    63 
       
    64 fun get_bn_fun_strs bn_funs =
       
    65   map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs
       
    66 *}
       
    67 
       
    68 
       
    69 ML {* 
       
    70 fun add_primrec_wrapper funs eqs lthy = 
       
    71   if null funs then (([], []), lthy)
       
    72   else 
       
    73    let 
       
    74      val eqs' = map (fn (_, eq) => (Attrib.empty_binding, eq)) eqs
       
    75      val funs' = map (fn (bn, ty, mx) => (bn, SOME ty, mx)) funs
       
    76      val ((funs'', eqs''), lthy') = Primrec.add_primrec funs' eqs' lthy
       
    77      val phi = ProofContext.export_morphism lthy' lthy
       
    78    in 
       
    79      ((map (Morphism.term phi) funs'', map (Morphism.thm phi) eqs''), lthy')
       
    80    end
       
    81 *}
       
    82 
       
    83 ML {*
       
    84 fun add_datatype_wrapper dt_names dts =
       
    85 let
       
    86   val conf = Datatype.default_config
       
    87 in
       
    88   Local_Theory.theory_result (Datatype.add_datatype conf dt_names dts)
       
    89 end
       
    90 *}
       
    91 
       
    92 
       
    93 text {* Infrastructure for adding "_raw" to types and terms *}
       
    94 
       
    95 ML {*
       
    96 fun add_raw s = s ^ "_raw"
       
    97 fun add_raws ss = map add_raw ss
       
    98 fun raw_bind bn = Binding.suffix_name "_raw" bn
       
    99 
       
   100 fun replace_str ss s = 
       
   101   case (AList.lookup (op=) ss s) of 
       
   102      SOME s' => s'
       
   103    | NONE => s
       
   104 
       
   105 fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts)
       
   106   | replace_typ ty_ss T = T  
       
   107 
       
   108 fun raw_dts ty_ss dts =
       
   109 let
       
   110   fun raw_dts_aux1 (bind, tys, mx) =
       
   111     (raw_bind bind, map (replace_typ ty_ss) tys, mx)
       
   112 
       
   113   fun raw_dts_aux2 (ty_args, bind, mx, constrs) =
       
   114     (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)
       
   115 in
       
   116   map raw_dts_aux2 dts
       
   117 end
       
   118 
       
   119 fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T)
       
   120   | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T)
       
   121   | replace_aterm trm_ss trm = trm
       
   122 
       
   123 fun replace_term trm_ss ty_ss trm =
       
   124   trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) 
       
   125 *}
       
   126 
       
   127 ML {*
       
   128 fun rawify_dts dt_names dts dts_env =
       
   129 let
       
   130   val raw_dts = raw_dts dts_env dts
       
   131   val raw_dt_names = add_raws dt_names
       
   132 in
       
   133   (raw_dt_names, raw_dts)
       
   134 end 
       
   135 *}
       
   136 
       
   137 ML {*
       
   138 fun rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs =
       
   139 let
       
   140   val bn_funs' = map (fn (bn, ty, mx) => 
       
   141     (raw_bind bn, replace_typ dts_env ty, mx)) bn_funs
       
   142   
       
   143   val bn_eqs' = map (fn (attr, trm) => 
       
   144     (attr, replace_term (cnstrs_env @ bn_fun_env) dts_env trm)) bn_eqs
       
   145 in
       
   146   (bn_funs', bn_eqs') 
       
   147 end 
       
   148 *}
       
   149 
       
   150 ML {* 
       
   151 fun rawify_bclauses dts_env cnstrs_env bn_fun_env bclauses =
       
   152 let
       
   153   fun rawify_bnds bnds = 
       
   154     map (apfst (Option.map (replace_term (cnstrs_env @ bn_fun_env) dts_env))) bnds
       
   155 
       
   156   fun rawify_bclause (BEmy n) = BEmy n
       
   157     | rawify_bclause (BLst (bnds, bdys)) = BLst (rawify_bnds bnds, bdys)
       
   158     | rawify_bclause (BSet (bnds, bdys)) = BSet (rawify_bnds bnds, bdys)
       
   159     | rawify_bclause (BRes (bnds, bdys)) = BRes (rawify_bnds bnds, bdys)
       
   160 in
       
   161   map (map (map rawify_bclause)) bclauses
       
   162 end
       
   163 *}
       
   164 
       
   165 (* strip_bn_fun takes a bn function defined by the user as a union or
       
   166    append of elements and returns those elements together with bn functions
       
   167    applied *)
       
   168 ML {*
       
   169 fun strip_bn_fun t =
       
   170   case t of
       
   171     Const (@{const_name sup}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
       
   172   | Const (@{const_name append}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
       
   173   | Const (@{const_name insert}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
       
   174       (i, NONE) :: strip_bn_fun y
       
   175   | Const (@{const_name Cons}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
       
   176       (i, NONE) :: strip_bn_fun y
       
   177   | Const (@{const_name bot}, _) => []
       
   178   | Const (@{const_name Nil}, _) => []
       
   179   | (f as Free _) $ Bound i => [(i, SOME f)]
       
   180   | _ => error ("Unsupported binding function: " ^ (PolyML.makestring t))
       
   181 *}
       
   182 
       
   183 ML {*
       
   184 fun find [] _ = error ("cannot find element")
       
   185   | find ((x, z)::xs) y = if (Long_Name.base_name x) = y then z else find xs y
       
   186 *}
       
   187 
       
   188 ML {*
       
   189 fun prep_bn lthy dt_names dts eqs = 
       
   190 let
       
   191   fun aux eq = 
       
   192   let
       
   193     val (lhs, rhs) = eq
       
   194       |> strip_qnt_body "all" 
       
   195       |> HOLogic.dest_Trueprop
       
   196       |> HOLogic.dest_eq
       
   197     val (bn_fun, [cnstr]) = strip_comb lhs
       
   198     val (_, ty) = dest_Free bn_fun
       
   199     val (ty_name, _) = dest_Type (domain_type ty)
       
   200     val dt_index = find_index (fn x => x = ty_name) dt_names
       
   201     val (cnstr_head, cnstr_args) = strip_comb cnstr
       
   202     val rhs_elements = strip_bn_fun rhs
       
   203     val included = map (apfst (fn i => length (cnstr_args) - i - 1)) rhs_elements
       
   204   in
       
   205     (dt_index, (bn_fun, (cnstr_head, included)))
       
   206   end
       
   207   fun aux2 eq = 
       
   208   let
       
   209     val (lhs, rhs) = eq
       
   210       |> strip_qnt_body "all" 
       
   211       |> HOLogic.dest_Trueprop
       
   212       |> HOLogic.dest_eq
       
   213     val (bn_fun, [cnstr]) = strip_comb lhs
       
   214     val (_, ty) = dest_Free bn_fun
       
   215     val (ty_name, _) = dest_Type (domain_type ty)
       
   216     val dt_index = find_index (fn x => x = ty_name) dt_names
       
   217     val (cnstr_head, cnstr_args) = strip_comb cnstr
       
   218     val rhs_elements = strip_bn_fun rhs
       
   219     val included = map (apfst (fn i => length (cnstr_args) - i - 1)) rhs_elements
       
   220   in
       
   221     (bn_fun, dt_index, (cnstr_head, included))
       
   222   end 
       
   223   fun order dts i ts = 
       
   224   let
       
   225     val dt = nth dts i
       
   226     val cts = map (fn (x, _, _) => Binding.name_of x) ((fn (_, _, _, x) => x) dt)
       
   227     val ts' = map (fn (x, y) => (fst (dest_Const x), y)) ts
       
   228   in
       
   229     map (find ts') cts
       
   230   end
       
   231 
       
   232   val unordered = AList.group (op=) (map aux eqs)
       
   233   val unordered' = map (fn (x, y) =>  (x, AList.group (op=) y)) unordered
       
   234   val ordered = map (fn (x, y) => (x, map (fn (v, z) => (v, order dts x z)) y)) unordered' 
       
   235   val ordered' = flat (map (fn (ith, l) => map (fn (bn, data) => (bn, ith, data)) l) ordered)
       
   236 
       
   237   val _ = tracing ("eqs\n" ^ cat_lines (map (Syntax.string_of_term lthy) eqs))
       
   238   val _ = tracing ("map eqs\n" ^ @{make_string} (map aux2 eqs))
       
   239   val _ = tracing ("ordered'\n" ^ @{make_string} ordered')
       
   240 in
       
   241   ordered'
       
   242 end
       
   243 *}
       
   244 
       
   245 
       
   246 ML {*
       
   247 fun raw_nominal_decls dts bn_funs bn_eqs binds lthy =
       
   248 let
       
   249   val thy = ProofContext.theory_of lthy
       
   250   val thy_name = Context.theory_name thy
       
   251 
       
   252   val dt_names = map (fn (_, s, _, _) => Binding.name_of s) dts
       
   253   val dt_full_names = map (Long_Name.qualify thy_name) dt_names 
       
   254   val dt_full_names' = add_raws dt_full_names
       
   255   val dts_env = dt_full_names ~~ dt_full_names'
       
   256 
       
   257   val cnstrs = get_cnstr_strs dts
       
   258   val cnstrs_ty = get_typed_cnstrs dts
       
   259   val cnstrs_full_names = map (Long_Name.qualify thy_name) cnstrs
       
   260   val cnstrs_full_names' = map (fn (x, y) => Long_Name.qualify thy_name 
       
   261     (Long_Name.qualify (add_raw x) (add_raw y))) cnstrs_ty
       
   262   val cnstrs_env = cnstrs_full_names ~~ cnstrs_full_names'
       
   263 
       
   264   val bn_fun_strs = get_bn_fun_strs bn_funs
       
   265   val bn_fun_strs' = add_raws bn_fun_strs
       
   266   val bn_fun_env = bn_fun_strs ~~ bn_fun_strs'
       
   267   val bn_fun_full_env = map (pairself (Long_Name.qualify thy_name)) 
       
   268     (bn_fun_strs ~~ bn_fun_strs')
       
   269   
       
   270   val (raw_dt_names, raw_dts) = rawify_dts dt_names dts dts_env
       
   271 
       
   272   val (raw_bn_funs, raw_bn_eqs) = rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs 
       
   273    
       
   274   val raw_bclauses = rawify_bclauses dts_env cnstrs_env bn_fun_full_env binds 
       
   275 
       
   276   val raw_bns = prep_bn lthy dt_full_names' raw_dts (map snd raw_bn_eqs)
       
   277 in
       
   278   lthy 
       
   279   |> add_datatype_wrapper raw_dt_names raw_dts 
       
   280   ||>> add_primrec_wrapper raw_bn_funs raw_bn_eqs
       
   281   ||>> pair raw_bclauses
       
   282   ||>> pair raw_bns
       
   283 end
       
   284 *}
       
   285 
       
   286 lemma equivp_hack: "equivp x"
       
   287 sorry
       
   288 ML {*
       
   289 fun equivp_hack ctxt rel =
       
   290 let
       
   291   val thy = ProofContext.theory_of ctxt
       
   292   val ty = domain_type (fastype_of rel)
       
   293   val cty = ctyp_of thy ty
       
   294   val ct = cterm_of thy rel
       
   295 in
       
   296   Drule.instantiate' [SOME cty] [SOME ct] @{thm equivp_hack}
       
   297 end
       
   298 *}
       
   299 
       
   300 ML {* val cheat_equivp = Unsynchronized.ref false *}
       
   301 ML {* val cheat_fv_rsp = Unsynchronized.ref false *}
       
   302 ML {* val cheat_alpha_bn_rsp = Unsynchronized.ref false *}
       
   303 ML {* val cheat_supp_eq = Unsynchronized.ref false *}
       
   304 
       
   305 ML {*
       
   306 fun remove_loop t =
       
   307   let val _ = HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of t)) in t end
       
   308   handle TERM _ => @{thm eqTrueI} OF [t]
       
   309 *}
       
   310 
       
   311 text {* 
       
   312   nominal_datatype2 does the following things in order:
       
   313 
       
   314 Parser.thy/raw_nominal_decls
       
   315   1) define the raw datatype
       
   316   2) define the raw binding functions 
       
   317 
       
   318 Perm.thy/define_raw_perms
       
   319   3) define permutations of the raw datatype and show that the raw type is 
       
   320      in the pt typeclass
       
   321       
       
   322 Lift.thy/define_fv_alpha_export, Fv.thy/define_fv & define_alpha
       
   323   4) define fv and fv_bn
       
   324   5) define alpha and alpha_bn
       
   325 
       
   326 Perm.thy/distinct_rel
       
   327   6) prove alpha_distincts (C1 x \<notsimeq> C2 y ...)             (Proof by cases; simp)
       
   328 
       
   329 Tacs.thy/build_rel_inj
       
   330   6) prove alpha_eq_iff    (C1 x = C2 y \<leftrightarrow> P x y ...)
       
   331      (left-to-right by intro rule, right-to-left by cases; simp)
       
   332 Equivp.thy/prove_eqvt
       
   333   7) prove bn_eqvt (common induction on the raw datatype)
       
   334   8) prove fv_eqvt (common induction on the raw datatype with help of above)
       
   335 Rsp.thy/build_alpha_eqvts
       
   336   9) prove alpha_eqvt and alpha_bn_eqvt
       
   337      (common alpha-induction, unfolding alpha_gen, permute of #* and =)
       
   338 Equivp.thy/build_alpha_refl & Equivp.thy/build_equivps
       
   339  10) prove that alpha and alpha_bn are equivalence relations
       
   340      (common induction and application of 'compose' lemmas)
       
   341 Lift.thy/define_quotient_types
       
   342  11) define quotient types
       
   343 Rsp.thy/build_fvbv_rsps
       
   344  12) prove bn respects     (common induction and simp with alpha_gen)
       
   345 Rsp.thy/prove_const_rsp
       
   346  13) prove fv respects     (common induction and simp with alpha_gen)
       
   347  14) prove permute respects    (unfolds to alpha_eqvt)
       
   348 Rsp.thy/prove_alpha_bn_rsp
       
   349  15) prove alpha_bn respects
       
   350      (alpha_induct then cases then sym and trans of the relations)
       
   351 Rsp.thy/prove_alpha_alphabn
       
   352  16) show that alpha implies alpha_bn (by unduction, needed in following step)
       
   353 Rsp.thy/prove_const_rsp
       
   354  17) prove respects for all datatype constructors
       
   355      (unfold eq_iff and alpha_gen; introduce zero permutations; simp)
       
   356 Perm.thy/quotient_lift_consts_export
       
   357  18) define lifted constructors, fv, bn, alpha_bn, permutations
       
   358 Perm.thy/define_lifted_perms
       
   359  19) lift permutation zero and add properties to show that quotient type is in the pt typeclass
       
   360 Lift.thy/lift_thm
       
   361  20) lift permutation simplifications
       
   362  21) lift induction
       
   363  22) lift fv
       
   364  23) lift bn
       
   365  24) lift eq_iff
       
   366  25) lift alpha_distincts
       
   367  26) lift fv and bn eqvts
       
   368 Equivp.thy/prove_supports
       
   369  27) prove that union of arguments supports constructors
       
   370 Equivp.thy/prove_fs
       
   371  28) show that the lifted type is in fs typeclass     (* by q_induct, supports *)
       
   372 Equivp.thy/supp_eq
       
   373  29) prove supp = fv
       
   374 *}
       
   375 
       
   376 ML {*
       
   377 (* for testing porposes - to exit the procedure early *)
       
   378 exception TEST of Proof.context
       
   379 
       
   380 val (STEPS, STEPS_setup) = Attrib.config_int "STEPS" (K 10);
       
   381 
       
   382 fun get_STEPS ctxt = Config.get ctxt STEPS
       
   383 *}
       
   384 
       
   385 setup STEPS_setup
       
   386 
       
   387 
       
   388 ML {*
       
   389 fun nominal_datatype2 dts bn_funs bn_eqs bclauses lthy =
       
   390 let
       
   391   (* definition of the raw datatype and raw bn-functions *)
       
   392   val ((((raw_dt_names, (raw_bn_funs, raw_bn_eqs)), raw_bclauses), raw_bns), lthy1) =
       
   393     if get_STEPS lthy > 1 then raw_nominal_decls dts bn_funs bn_eqs bclauses lthy
       
   394     else raise TEST lthy
       
   395 
       
   396   val dtinfo = Datatype.the_info (ProofContext.theory_of lthy1) (hd raw_dt_names);
       
   397   val {descr, sorts, ...} = dtinfo;
       
   398   val raw_tys = map (fn (i, _) => nth_dtyp descr sorts i) descr;
       
   399 
       
   400   val induct_thm = #induct dtinfo;
       
   401 
       
   402   (* definitions of raw permutations *)
       
   403   val ((raw_perm_def, raw_perm_simps, perms), lthy2) =
       
   404     if get_STEPS lthy > 2 
       
   405     then Local_Theory.theory_result (define_raw_perms descr sorts induct_thm (length dts)) lthy1
       
   406     else raise TEST lthy1
       
   407 
       
   408   (* definition of raw fv_functions *)
       
   409   val morphism_2_0 = ProofContext.export_morphism lthy2 lthy
       
   410   fun export_fun f (t, n , l) = (f t, n, map (map (apsnd (Option.map f))) l);
       
   411   val bn_funs_decls = map (export_fun (Morphism.term morphism_2_0)) raw_bns;
       
   412  
       
   413   val thy = Local_Theory.exit_global lthy2;
       
   414   val thy_name = Context.theory_name thy
       
   415 
       
   416   val lthy3 = Theory_Target.init NONE thy;
       
   417   val raw_bn_funs = map (fn (f, _, _) => f) bn_funs_decls;
       
   418 
       
   419   val _ = tracing ("raw_bns\n" ^ @{make_string} raw_bns)
       
   420   val _ = tracing ("bn_funs\n" ^ @{make_string} bn_funs_decls)
       
   421    
       
   422   val ((fv, fvbn), fv_def, lthy3a) = 
       
   423     if get_STEPS lthy > 3 
       
   424     then define_raw_fv descr sorts bn_funs_decls raw_bclauses lthy3
       
   425     else raise TEST lthy3
       
   426   
       
   427 in
       
   428   (0, lthy3a)
       
   429 end handle TEST ctxt => (0, ctxt)
       
   430 *}
       
   431 
       
   432 section {* Preparing and parsing of the specification *}
       
   433 
       
   434 ML {* 
       
   435 (* parsing the datatypes and declaring *)
       
   436 (* constructors in the local theory    *)
       
   437 fun prepare_dts dt_strs lthy = 
       
   438 let
       
   439   val thy = ProofContext.theory_of lthy
       
   440   
       
   441   fun mk_type full_tname tvrs =
       
   442     Type (full_tname, map (fn a => TVar ((a, 0), [])) tvrs)
       
   443 
       
   444   fun prep_cnstr full_tname tvs (cname, anno_tys, mx, _) =
       
   445   let
       
   446     val tys = map (Syntax.read_typ lthy o snd) anno_tys
       
   447     val ty = mk_type full_tname tvs
       
   448   in
       
   449     ((cname, tys ---> ty, mx), (cname, tys, mx))
       
   450   end
       
   451   
       
   452   fun prep_dt (tvs, tname, mx, cnstrs) = 
       
   453   let
       
   454     val full_tname = Sign.full_name thy tname
       
   455     val (cnstrs', cnstrs'') = 
       
   456       split_list (map (prep_cnstr full_tname tvs) cnstrs)
       
   457   in
       
   458     (cnstrs', (tvs, tname, mx, cnstrs''))
       
   459   end 
       
   460 
       
   461   val (cnstrs, dts) = split_list (map prep_dt dt_strs)
       
   462 in
       
   463   lthy
       
   464   |> Local_Theory.theory (Sign.add_consts_i (flat cnstrs))
       
   465   |> pair dts
       
   466 end
       
   467 *}
       
   468 
       
   469 ML {*
       
   470 (* parsing the binding function specification and *)
       
   471 (* declaring the functions in the local theory    *)
       
   472 fun prepare_bn_funs bn_fun_strs bn_eq_strs lthy =
       
   473 let
       
   474   val ((bn_funs, bn_eqs), _) = 
       
   475     Specification.read_spec bn_fun_strs bn_eq_strs lthy
       
   476 
       
   477   fun prep_bn_fun ((bn, T), mx) = (bn, T, mx) 
       
   478   
       
   479   val bn_funs' = map prep_bn_fun bn_funs
       
   480 in
       
   481   lthy
       
   482   |> Local_Theory.theory (Sign.add_consts_i bn_funs')
       
   483   |> pair (bn_funs', bn_eqs) 
       
   484 end 
       
   485 *}
       
   486 
       
   487 text {* associates every SOME with the index in the list; drops NONEs *}
       
   488 ML {*
       
   489 fun indexify xs =
       
   490 let
       
   491   fun mapp _ [] = []
       
   492     | mapp i (NONE :: xs) = mapp (i + 1) xs
       
   493     | mapp i (SOME x :: xs) = (x, i) :: mapp (i + 1) xs
       
   494 in 
       
   495   mapp 0 xs 
       
   496 end
       
   497 
       
   498 fun index_lookup xs x =
       
   499   case AList.lookup (op=) xs x of
       
   500     SOME x => x
       
   501   | NONE => error ("Cannot find " ^ x ^ " as argument annotation.");
       
   502 *}
       
   503 
       
   504 ML {*
       
   505 fun prepare_bclauses dt_strs lthy = 
       
   506 let
       
   507   val annos_bclauses =
       
   508     get_cnstrs dt_strs
       
   509     |> map (map (fn (_, antys, _, bns) => (map fst antys, bns)))
       
   510 
       
   511   fun prep_binder env bn_str =
       
   512     case (Syntax.read_term lthy bn_str) of
       
   513       Free (x, _) => (NONE, index_lookup env x)
       
   514     | Const (a, T) $ Free (x, _) => (SOME (Const (a, T)), index_lookup env x)
       
   515     | _ => error ("The term " ^ bn_str ^ " is not allowed as binding function.")
       
   516  
       
   517   fun prep_body env bn_str = index_lookup env bn_str
       
   518 
       
   519   fun prep_mode "bind"     = BLst 
       
   520     | prep_mode "bind_set" = BSet 
       
   521     | prep_mode "bind_res" = BRes 
       
   522 
       
   523   fun prep_bclause env (mode, binders, bodies) = 
       
   524   let
       
   525     val binders' = map (prep_binder env) binders
       
   526     val bodies' = map (prep_body env) bodies
       
   527   in  
       
   528     prep_mode mode (binders', bodies')
       
   529   end
       
   530 
       
   531   fun prep_bclauses (annos, bclause_strs) = 
       
   532   let
       
   533     val env = indexify annos (* for every label, associate the index *)
       
   534   in
       
   535     map (prep_bclause env) bclause_strs
       
   536   end
       
   537 in
       
   538   map (map prep_bclauses) annos_bclauses
       
   539 end
       
   540 *}
       
   541 
       
   542 text {* 
       
   543   adds an empty binding clause for every argument
       
   544   that is not already part of a binding clause
       
   545 *}
       
   546 
       
   547 ML {*
       
   548 fun included i bcs = 
       
   549 let
       
   550   fun incl (BEmy j) = i = j
       
   551     | incl (BLst (bns, bds)) = (member (op =) (map snd bns) i) orelse (member (op =) bds i)
       
   552     | incl (BSet (bns, bds)) = (member (op =) (map snd bns) i) orelse (member (op =) bds i)
       
   553     | incl (BRes (bns, bds)) = (member (op =) (map snd bns) i) orelse (member (op =) bds i)
       
   554 in
       
   555   exists incl bcs 
       
   556 end
       
   557 *}
       
   558 
       
   559 ML {* 
       
   560 fun complete dt_strs bclauses = 
       
   561 let
       
   562   val args = 
       
   563     get_cnstrs dt_strs
       
   564     |> map (map (fn (_, antys, _, _) => length antys))
       
   565 
       
   566   fun complt n bcs = 
       
   567   let
       
   568     fun add bcs i = (if included i bcs then [] else [BEmy i]) 
       
   569   in
       
   570     bcs @ (flat (map_range (add bcs) n))
       
   571   end
       
   572 in
       
   573   map2 (map2 complt) args bclauses
       
   574 end
       
   575 *}
       
   576 
       
   577 ML {*
       
   578 fun nominal_datatype2_cmd (dt_strs, bn_fun_strs, bn_eq_strs) lthy = 
       
   579 let
       
   580   fun prep_typ (tvs, tname, mx, _) = (tname, length tvs, mx)
       
   581   val lthy0 = 
       
   582     Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) lthy
       
   583   val (dts, lthy1) = prepare_dts dt_strs lthy0
       
   584   val ((bn_funs, bn_eqs), lthy2) = prepare_bn_funs bn_fun_strs bn_eq_strs lthy1
       
   585   val bclauses = prepare_bclauses dt_strs lthy2
       
   586   val bclauses' = complete dt_strs bclauses 
       
   587 in
       
   588   nominal_datatype2 dts bn_funs bn_eqs bclauses' lthy |> snd
       
   589 end
       
   590 
       
   591 
       
   592 (* Command Keyword *)
       
   593 
       
   594 val _ = OuterSyntax.local_theory "nominal_datatype" "test" OuterKeyword.thy_decl
       
   595   (main_parser >> nominal_datatype2_cmd)
       
   596 *}
       
   597 
       
   598 (*
       
   599 atom_decl name
       
   600 
       
   601 nominal_datatype lam =
       
   602   Var name
       
   603 | App lam lam
       
   604 | Lam x::name t::lam  bind_set x in t
       
   605 | Let p::pt t::lam bind_set "bn p" in t
       
   606 and pt =
       
   607   P1 name
       
   608 | P2 pt pt
       
   609 binder
       
   610  bn::"pt \<Rightarrow> atom set"
       
   611 where
       
   612   "bn (P1 x) = {atom x}"
       
   613 | "bn (P2 p1 p2) = bn p1 \<union> bn p2"
       
   614 
       
   615 find_theorems Var_raw
       
   616 
       
   617 
       
   618 
       
   619 thm lam_pt.bn
       
   620 thm lam_pt.fv[simplified lam_pt.supp(1-2)]
       
   621 thm lam_pt.eq_iff
       
   622 thm lam_pt.induct
       
   623 thm lam_pt.perm
       
   624 
       
   625 nominal_datatype exp =
       
   626   EVar name
       
   627 | EUnit
       
   628 | EPair q1::exp q2::exp
       
   629 | ELetRec l::lrbs e::exp     bind "b_lrbs l" in e l
       
   630 
       
   631 and fnclause =
       
   632   K x::name p::pat f::exp    bind_res "b_pat p" in f
       
   633 
       
   634 and fnclauses =
       
   635   S fnclause
       
   636 | ORs fnclause fnclauses
       
   637 
       
   638 and lrb =
       
   639   Clause fnclauses
       
   640 
       
   641 and lrbs =
       
   642   Single lrb
       
   643 | More lrb lrbs
       
   644 
       
   645 and pat =
       
   646   PVar name
       
   647 | PUnit
       
   648 | PPair pat pat
       
   649 
       
   650 binder
       
   651   b_lrbs      :: "lrbs \<Rightarrow> atom list" and
       
   652   b_pat       :: "pat \<Rightarrow> atom set" and
       
   653   b_fnclauses :: "fnclauses \<Rightarrow> atom list" and
       
   654   b_fnclause  :: "fnclause \<Rightarrow> atom list" and
       
   655   b_lrb       :: "lrb \<Rightarrow> atom list"
       
   656 
       
   657 where
       
   658   "b_lrbs (Single l) = b_lrb l"
       
   659 | "b_lrbs (More l ls) = append (b_lrb l) (b_lrbs ls)"
       
   660 | "b_pat (PVar x) = {atom x}"
       
   661 | "b_pat (PUnit) = {}"
       
   662 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
       
   663 | "b_fnclauses (S fc) = (b_fnclause fc)"
       
   664 | "b_fnclauses (ORs fc fcs) = append (b_fnclause fc) (b_fnclauses fcs)"
       
   665 | "b_lrb (Clause fcs) = (b_fnclauses fcs)"
       
   666 | "b_fnclause (K x pat exp) = [atom x]"
       
   667 
       
   668 thm exp_fnclause_fnclauses_lrb_lrbs_pat.bn
       
   669 thm exp_fnclause_fnclauses_lrb_lrbs_pat.fv
       
   670 thm exp_fnclause_fnclauses_lrb_lrbs_pat.eq_iff
       
   671 thm exp_fnclause_fnclauses_lrb_lrbs_pat.induct
       
   672 thm exp_fnclause_fnclauses_lrb_lrbs_pat.perm
       
   673 
       
   674 nominal_datatype ty =
       
   675   Vr "name"
       
   676 | Fn "ty" "ty"
       
   677 and tys =
       
   678   Al xs::"name fset" t::"ty" bind_res xs in t
       
   679 
       
   680 thm ty_tys.fv[simplified ty_tys.supp]
       
   681 thm ty_tys.eq_iff
       
   682 
       
   683 *)
       
   684 
       
   685 (* some further tests *)
       
   686 
       
   687 (*
       
   688 nominal_datatype ty2 =
       
   689   Vr2 "name"
       
   690 | Fn2 "ty2" "ty2"
       
   691 
       
   692 nominal_datatype tys2 =
       
   693   All2 xs::"name fset" ty::"ty2" bind_res xs in ty
       
   694 
       
   695 nominal_datatype lam2 =
       
   696   Var2 "name"
       
   697 | App2 "lam2" "lam2 list"
       
   698 | Lam2 x::"name" t::"lam2" bind x in t
       
   699 *)
       
   700 
       
   701 
       
   702 
       
   703 end
       
   704 
       
   705 
       
   706