Nominal/NewParser.thy
changeset 1941 d33781f9d2c7
child 1942 d3b89f6c086a
equal deleted inserted replaced
1940:0913f697fe73 1941:d33781f9d2c7
       
     1 theory NewParser
       
     2 imports "../Nominal-General/Nominal2_Atoms" 
       
     3         "../Nominal-General/Nominal2_Eqvt" 
       
     4         "../Nominal-General/Nominal2_Supp" 
       
     5         "Perm" "Equivp" "Rsp" "Lift"
       
     6 begin
       
     7 
       
     8 section{* Interface for nominal_datatype *}
       
     9 
       
    10 text {*
       
    11 
       
    12 Nominal-Datatype-part:
       
    13 
       
    14 
       
    15 1nd Arg: (string list * binding * mixfix * (binding * typ list * mixfix) list) list
       
    16          ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^   ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
       
    17                type(s) to be defined             constructors list
       
    18                (ty args, name, syn)              (name, typs, syn)
       
    19 
       
    20 Binder-Function-part:
       
    21 
       
    22 2rd Arg: (binding * typ option * mixfix) list 
       
    23          ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^    
       
    24             binding function(s)           
       
    25               to be defined               
       
    26             (name, type, syn)             
       
    27 
       
    28 3th Arg:  term list 
       
    29           ^^^^^^^^^
       
    30           the equations of the binding functions
       
    31           (Trueprop equations)
       
    32 *}
       
    33 
       
    34 text {*****************************************************}
       
    35 ML {* OuterParse.triple2 *}
       
    36 
       
    37 ML {* 
       
    38 (* nominal datatype parser *)
       
    39 local
       
    40   structure P = OuterParse;
       
    41   structure S = Scan
       
    42 
       
    43   fun triple1 ((x, y), z) = (x, y, z)
       
    44   fun triple2 (x, (y, z)) = (x, y, z)
       
    45   fun tuple ((x, y, z), u) = (x, y, z, u)
       
    46   fun tswap (((x, y), z), u) = (x, y, u, z)
       
    47 in
       
    48 
       
    49 val _ = OuterKeyword.keyword "bind"
       
    50 val _ = OuterKeyword.keyword "bind_set"
       
    51 val _ = OuterKeyword.keyword "bind_res"
       
    52 
       
    53 val anno_typ = Scan.option (P.name --| P.$$$ "::") -- P.typ
       
    54 
       
    55 val bind_mode = P.$$$ "bind" || P.$$$ "bind_set" || P.$$$ "bind_res"
       
    56 
       
    57 val bind_clauses = 
       
    58   P.enum "," (bind_mode -- S.repeat1 P.term -- (P.$$$ "in" |-- S.repeat1 P.name) >> triple1)
       
    59 
       
    60 val cnstr_parser =
       
    61   P.binding -- Scan.repeat anno_typ
       
    62 
       
    63 (* datatype parser *)
       
    64 val dt_parser =
       
    65   (P.type_args -- P.binding -- P.opt_mixfix >> triple1) -- 
       
    66     (P.$$$ "=" |-- P.enum1 "|" (cnstr_parser -- bind_clauses -- P.opt_mixfix >> tswap)) >> tuple
       
    67 
       
    68 (* binding function parser *)
       
    69 val bnfun_parser = 
       
    70   Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[])
       
    71 
       
    72 (* main parser *)
       
    73 val main_parser =
       
    74   (P.and_list1 dt_parser) -- bnfun_parser >> triple2
       
    75 
       
    76 end
       
    77 *}
       
    78 
       
    79 ML {*
       
    80 datatype bmodes =
       
    81    BEmy of int
       
    82 |  BLst of ((term option * int) list) * (int list)
       
    83 |  BSet of ((term option * int) list) * (int list)
       
    84 |  BRes of ((term option * int) list) * (int list)
       
    85 *}
       
    86 
       
    87 ML {*
       
    88 fun get_cnstrs dts =
       
    89   map (fn (_, _, _, constrs) => constrs) dts
       
    90 
       
    91 fun get_typed_cnstrs dts =
       
    92   flat (map (fn (_, bn, _, constrs) => 
       
    93    (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts)
       
    94 
       
    95 fun get_cnstr_strs dts =
       
    96   map (fn (bn, _, _) => Binding.name_of bn) (flat (get_cnstrs dts))
       
    97 
       
    98 fun get_bn_fun_strs bn_funs =
       
    99   map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs
       
   100 *}
       
   101 
       
   102 ML {* 
       
   103 fun add_primrec_wrapper funs eqs lthy = 
       
   104   if null funs then (([], []), lthy)
       
   105   else 
       
   106    let 
       
   107      val eqs' = map (fn (_, eq) => (Attrib.empty_binding, eq)) eqs
       
   108      val funs' = map (fn (bn, ty, mx) => (bn, SOME ty, mx)) funs
       
   109    in 
       
   110      Primrec.add_primrec funs' eqs' lthy
       
   111    end
       
   112 *}
       
   113 
       
   114 ML {*
       
   115 fun add_datatype_wrapper dt_names dts =
       
   116 let
       
   117   val conf = Datatype.default_config
       
   118 in
       
   119   Local_Theory.theory_result (Datatype.add_datatype conf dt_names dts)
       
   120 end
       
   121 *}
       
   122 
       
   123 ML {*
       
   124 fun add_raw s = s ^ "_raw"
       
   125 fun add_raws ss = map add_raw ss
       
   126 fun raw_bind bn = Binding.suffix_name "_raw" bn
       
   127 
       
   128 fun replace_str ss s = 
       
   129   case (AList.lookup (op=) ss s) of 
       
   130      SOME s' => s'
       
   131    | NONE => s
       
   132 
       
   133 fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts)
       
   134   | replace_typ ty_ss T = T  
       
   135 
       
   136 fun raw_dts ty_ss dts =
       
   137 let
       
   138   fun raw_dts_aux1 (bind, tys, mx) =
       
   139     (raw_bind bind, map (replace_typ ty_ss) tys, mx)
       
   140 
       
   141   fun raw_dts_aux2 (ty_args, bind, mx, constrs) =
       
   142     (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)
       
   143 in
       
   144   map raw_dts_aux2 dts
       
   145 end
       
   146 
       
   147 fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T)
       
   148   | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T)
       
   149   | replace_aterm trm_ss trm = trm
       
   150 
       
   151 fun replace_term trm_ss ty_ss trm =
       
   152   trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) 
       
   153 *}
       
   154 
       
   155 ML {*
       
   156 fun rawify_dts dt_names dts dts_env =
       
   157 let
       
   158   val raw_dts = raw_dts dts_env dts
       
   159   val raw_dt_names = add_raws dt_names
       
   160 in
       
   161   (raw_dt_names, raw_dts)
       
   162 end 
       
   163 *}
       
   164 
       
   165 ML {*
       
   166 fun rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs =
       
   167 let
       
   168   val bn_funs' = map (fn (bn, ty, mx) => 
       
   169     (raw_bind bn, replace_typ dts_env ty, mx)) bn_funs
       
   170   
       
   171   val bn_eqs' = map (fn (attr, trm) => 
       
   172     (attr, replace_term (cnstrs_env @ bn_fun_env) dts_env trm)) bn_eqs
       
   173 in
       
   174   (bn_funs', bn_eqs') 
       
   175 end 
       
   176 *}
       
   177 
       
   178 ML {* 
       
   179 fun rawify_bclauses dts_env cnstrs_env bn_fun_env bclauses =
       
   180 let
       
   181   fun rawify_bnds bnds = 
       
   182     map (apfst (Option.map (replace_term (cnstrs_env @ bn_fun_env) dts_env))) bnds
       
   183 
       
   184   fun rawify_bclause (BEmy n) = BEmy n
       
   185     | rawify_bclause (BLst (bnds, bdys)) = BLst (rawify_bnds bnds, bdys)
       
   186     | rawify_bclause (BSet (bnds, bdys)) = BSet (rawify_bnds bnds, bdys)
       
   187     | rawify_bclause (BRes (bnds, bdys)) = BRes (rawify_bnds bnds, bdys)
       
   188 
       
   189 in
       
   190   map (map (map rawify_bclause)) bclauses
       
   191 end
       
   192 *}
       
   193 
       
   194 text {* What does the prep_bn code do ? *}
       
   195 
       
   196 ML {*
       
   197 fun strip_bn_fun t =
       
   198   case t of
       
   199     Const (@{const_name sup}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
       
   200   | Const (@{const_name append}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
       
   201   | Const (@{const_name insert}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
       
   202       (i, NONE) :: strip_bn_fun y
       
   203   | Const (@{const_name Cons}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
       
   204       (i, NONE) :: strip_bn_fun y
       
   205   | Const (@{const_name bot}, _) => []
       
   206   | Const (@{const_name Nil}, _) => []
       
   207   | (f as Free _) $ Bound i => [(i, SOME f)]
       
   208   | _ => error ("Unsupported binding function: " ^ (PolyML.makestring t))
       
   209 *}
       
   210 
       
   211 ML {*
       
   212 fun find [] _ = error ("cannot find element")
       
   213   | find ((x, z)::xs) y = if (Long_Name.base_name x) = y then z else find xs y
       
   214 *}
       
   215 
       
   216 ML {*
       
   217 fun prep_bn dt_names dts eqs = 
       
   218 let
       
   219   fun aux eq = 
       
   220   let
       
   221     val (lhs, rhs) = eq
       
   222       |> strip_qnt_body "all" 
       
   223       |> HOLogic.dest_Trueprop
       
   224       |> HOLogic.dest_eq
       
   225     val (bn_fun, [cnstr]) = strip_comb lhs
       
   226     val (_, ty) = dest_Free bn_fun
       
   227     val (ty_name, _) = dest_Type (domain_type ty)
       
   228     val dt_index = find_index (fn x => x = ty_name) dt_names
       
   229     val (cnstr_head, cnstr_args) = strip_comb cnstr
       
   230     val rhs_elements = strip_bn_fun rhs
       
   231     val included = map (apfst (fn i => length (cnstr_args) - i - 1)) rhs_elements
       
   232   in
       
   233     (dt_index, (bn_fun, (cnstr_head, included)))
       
   234   end 
       
   235   fun order dts i ts = 
       
   236   let
       
   237     val dt = nth dts i
       
   238     val cts = map (fn (x, _, _) => Binding.name_of x) ((fn (_, _, _, x) => x) dt)
       
   239     val ts' = map (fn (x, y) => (fst (dest_Const x), y)) ts
       
   240   in
       
   241     map (find ts') cts
       
   242   end
       
   243 
       
   244   val unordered = AList.group (op=) (map aux eqs)
       
   245   val unordered' = map (fn (x, y) =>  (x, AList.group (op=) y)) unordered
       
   246   val ordered = map (fn (x, y) => (x, map (fn (v, z) => (v, order dts x z)) y)) unordered' 
       
   247 in
       
   248   ordered
       
   249 end
       
   250 *}
       
   251 
       
   252 
       
   253 ML {*
       
   254 fun raw_nominal_decls dts bn_funs bn_eqs binds lthy =
       
   255 let
       
   256   val thy = ProofContext.theory_of lthy
       
   257   val thy_name = Context.theory_name thy
       
   258 
       
   259   val dt_names = map (fn (_, s, _, _) => Binding.name_of s) dts
       
   260   val dt_full_names = map (Long_Name.qualify thy_name) dt_names 
       
   261   val dt_full_names' = add_raws dt_full_names
       
   262   val dts_env = dt_full_names ~~ dt_full_names'
       
   263 
       
   264   val cnstrs = get_cnstr_strs dts
       
   265   val cnstrs_ty = get_typed_cnstrs dts
       
   266   val cnstrs_full_names = map (Long_Name.qualify thy_name) cnstrs
       
   267   val cnstrs_full_names' = map (fn (x, y) => Long_Name.qualify thy_name 
       
   268     (Long_Name.qualify (add_raw x) (add_raw y))) cnstrs_ty
       
   269   val cnstrs_env = cnstrs_full_names ~~ cnstrs_full_names'
       
   270 
       
   271   val bn_fun_strs = get_bn_fun_strs bn_funs
       
   272   val bn_fun_strs' = add_raws bn_fun_strs
       
   273   val bn_fun_env = bn_fun_strs ~~ bn_fun_strs'
       
   274   val bn_fun_full_env = map (pairself (Long_Name.qualify thy_name)) 
       
   275     (bn_fun_strs ~~ bn_fun_strs')
       
   276   
       
   277   val (raw_dt_names, raw_dts) = rawify_dts dt_names dts dts_env
       
   278 
       
   279   val (raw_bn_funs, raw_bn_eqs) = rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs 
       
   280    
       
   281   val raw_bclauses = rawify_bclauses dts_env cnstrs_env bn_fun_full_env binds 
       
   282 
       
   283   val raw_bns = prep_bn dt_full_names' raw_dts (map snd raw_bn_eqs)
       
   284 
       
   285 in
       
   286   lthy 
       
   287   |> add_datatype_wrapper raw_dt_names raw_dts 
       
   288   ||>> add_primrec_wrapper raw_bn_funs raw_bn_eqs
       
   289   ||>> pair raw_bclauses
       
   290   ||>> pair raw_bns
       
   291 end
       
   292 *}
       
   293 
       
   294 ML {*
       
   295 fun nominal_datatype2 dts bn_funs bn_eqs bclauses lthy =
       
   296 let
       
   297   val ((((raw_dt_names, (raw_bn_funs_loc, raw_bn_eqs_loc)), raw_bclauses), raw_bns), lthy2) =
       
   298     raw_nominal_decls dts bn_funs bn_eqs bclauses lthy
       
   299 
       
   300 in
       
   301   ((raw_dt_names, raw_bn_funs_loc, raw_bn_eqs_loc, raw_bclauses, raw_bns), lthy2)
       
   302 end
       
   303 *}
       
   304 
       
   305 section {* Preparing and parsing of the specification *}
       
   306 
       
   307 ML {* 
       
   308 (* parsing the datatypes and declaring *)
       
   309 (* constructors in the local theory    *)
       
   310 fun prepare_dts dt_strs lthy = 
       
   311 let
       
   312   val thy = ProofContext.theory_of lthy
       
   313   
       
   314   fun mk_type full_tname tvrs =
       
   315     Type (full_tname, map (fn a => TVar ((a, 0), [])) tvrs)
       
   316 
       
   317   fun prep_cnstr full_tname tvs (cname, anno_tys, mx, _) =
       
   318   let
       
   319     val tys = map (Syntax.read_typ lthy o snd) anno_tys
       
   320     val ty = mk_type full_tname tvs
       
   321   in
       
   322     ((cname, tys ---> ty, mx), (cname, tys, mx))
       
   323   end
       
   324   
       
   325   fun prep_dt (tvs, tname, mx, cnstrs) = 
       
   326   let
       
   327     val full_tname = Sign.full_name thy tname
       
   328     val (cnstrs', cnstrs'') = 
       
   329       split_list (map (prep_cnstr full_tname tvs) cnstrs)
       
   330   in
       
   331     (cnstrs', (tvs, tname, mx, cnstrs''))
       
   332   end 
       
   333 
       
   334   val (cnstrs, dts) = split_list (map prep_dt dt_strs)
       
   335 in
       
   336   lthy
       
   337   |> Local_Theory.theory (Sign.add_consts_i (flat cnstrs))
       
   338   |> pair dts
       
   339 end
       
   340 *}
       
   341 
       
   342 ML {*
       
   343 (* parsing the binding function specification and *)
       
   344 (* declaring the functions in the local theory    *)
       
   345 fun prepare_bn_funs bn_fun_strs bn_eq_strs lthy =
       
   346 let
       
   347   val ((bn_funs, bn_eqs), _) = 
       
   348     Specification.read_spec bn_fun_strs bn_eq_strs lthy
       
   349 
       
   350   fun prep_bn_fun ((bn, T), mx) = (bn, T, mx) 
       
   351   
       
   352   val bn_funs' = map prep_bn_fun bn_funs
       
   353 in
       
   354   lthy
       
   355   |> Local_Theory.theory (Sign.add_consts_i bn_funs')
       
   356   |> pair (bn_funs', bn_eqs) 
       
   357 end 
       
   358 *}
       
   359 
       
   360 text {* associates every SOME with the index in the list; drops NONEs *}
       
   361 ML {*
       
   362 fun indexify xs =
       
   363 let
       
   364   fun mapp _ [] = []
       
   365     | mapp i (NONE :: xs) = mapp (i + 1) xs
       
   366     | mapp i (SOME x :: xs) = (x, i) :: mapp (i + 1) xs
       
   367 in 
       
   368   mapp 0 xs 
       
   369 end
       
   370 
       
   371 fun index_lookup xs x =
       
   372   case AList.lookup (op=) xs x of
       
   373     SOME x => x
       
   374   | NONE => error ("Cannot find " ^ x ^ " as argument annotation.");
       
   375 *}
       
   376 
       
   377 ML {*
       
   378 fun prepare_bclauses dt_strs lthy = 
       
   379 let
       
   380   val annos_bclauses =
       
   381     get_cnstrs dt_strs
       
   382     |> map (map (fn (_, antys, _, bns) => (map fst antys, bns)))
       
   383 
       
   384   fun prep_binder env bn_str =
       
   385     case (Syntax.read_term lthy bn_str) of
       
   386       Free (x, _) => (NONE, index_lookup env x)
       
   387     | Const (a, T) $ Free (x, _) => (SOME (Const (a, T)), index_lookup env x)
       
   388     | _ => error ("The term " ^ bn_str ^ " is not allowed as binding function.")
       
   389  
       
   390   fun prep_body env bn_str = index_lookup env bn_str
       
   391 
       
   392   fun prep_mode "bind"     = BLst 
       
   393     | prep_mode "bind_set" = BSet 
       
   394     | prep_mode "bind_res" = BRes 
       
   395 
       
   396   fun prep_bclause env (mode, binders, bodies) = 
       
   397   let
       
   398     val binders' = map (prep_binder env) binders
       
   399     val bodies' = map (prep_body env) bodies
       
   400   in  
       
   401     prep_mode mode (binders', bodies')
       
   402   end
       
   403 
       
   404   fun prep_bclauses (annos, bclause_strs) = 
       
   405   let
       
   406     val env = indexify annos (* for every label, associate the index *)
       
   407   in
       
   408     map (prep_bclause env) bclause_strs
       
   409   end
       
   410 in
       
   411   map (map prep_bclauses) annos_bclauses
       
   412 end
       
   413 *}
       
   414 
       
   415 ML {*
       
   416 fun included i bcs = 
       
   417 let
       
   418   fun incl (BEmy j) = i = j
       
   419     | incl (BLst (bns, bds)) = (i mem (map snd bns)) orelse (i mem bds) 
       
   420     | incl (BSet (bns, bds)) = (i mem (map snd bns)) orelse (i mem bds) 
       
   421     | incl (BRes (bns, bds)) = (i mem (map snd bns)) orelse (i mem bds) 
       
   422 in
       
   423   exists incl bcs 
       
   424 end
       
   425 *}
       
   426 
       
   427 ML {* 
       
   428 fun complete dt_strs bclauses = 
       
   429 let
       
   430   val args = 
       
   431     get_cnstrs dt_strs
       
   432     |> map (map (fn (_, antys, _, _) => length antys))
       
   433 
       
   434   fun complt n bcs = 
       
   435   let
       
   436     fun add bcs i = (if included i bcs then [] else [BEmy i]) 
       
   437   in
       
   438     bcs @ (flat (map_range (add bcs) n))
       
   439   end
       
   440 in
       
   441   map2 (map2 complt) args bclauses
       
   442 end
       
   443 *}
       
   444 
       
   445 ML {*
       
   446 fun nominal_datatype2_cmd (dt_strs, bn_fun_strs, bn_eq_strs) lthy = 
       
   447 let
       
   448   fun prep_typ (tvs, tname, mx, _) = (tname, length tvs, mx)
       
   449 
       
   450   val lthy0 = 
       
   451     Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) lthy
       
   452   val (dts, lthy1) = 
       
   453     prepare_dts dt_strs lthy0
       
   454   val ((bn_funs, bn_eqs), lthy2) = 
       
   455     prepare_bn_funs bn_fun_strs bn_eq_strs lthy1
       
   456   val bclauses = prepare_bclauses dt_strs lthy2
       
   457   val bclauses' = complete dt_strs bclauses 
       
   458 in
       
   459   nominal_datatype2 dts bn_funs bn_eqs bclauses' lthy |> snd
       
   460 end
       
   461 
       
   462 
       
   463 (* Command Keyword *)
       
   464 
       
   465 val _ = OuterSyntax.local_theory "nominal_datatype" "test" OuterKeyword.thy_decl
       
   466   (main_parser >> nominal_datatype2_cmd)
       
   467 
       
   468 *}
       
   469 
       
   470 
       
   471 
       
   472 atom_decl name
       
   473 
       
   474 nominal_datatype exp =
       
   475   EVar name
       
   476 | EUnit
       
   477 | EPair q1::exp q2::exp      bind_set q1 q2 in q1 q2
       
   478 | ELetRec l::lrbs e::exp     bind "b_lrbs l" in e l
       
   479 
       
   480 and fnclause =
       
   481   K x::name p::pat f::exp    bind_res "b_pat p" in f 
       
   482 
       
   483 and fnclauses =
       
   484   S fnclause
       
   485 | ORs fnclause fnclauses
       
   486 
       
   487 and lrb =
       
   488   Clause fnclauses
       
   489 
       
   490 and lrbs =
       
   491   Single lrb
       
   492 | More lrb lrbs
       
   493 
       
   494 and pat =
       
   495   PVar name
       
   496 | PUnit
       
   497 | PPair pat pat
       
   498 
       
   499 binder
       
   500   b_lrbs      :: "lrbs \<Rightarrow> atom set" and
       
   501   b_pat       :: "pat \<Rightarrow> atom set" and
       
   502   b_fnclauses :: "fnclauses \<Rightarrow> atom set" and
       
   503   b_fnclause  :: "fnclause \<Rightarrow> atom set" and
       
   504   b_lrb       :: "lrb \<Rightarrow> atom set"
       
   505 
       
   506 where
       
   507   "b_lrbs (Single l) = b_lrb l"
       
   508 | "b_lrbs (More l ls) = b_lrb l \<union> b_lrbs ls"
       
   509 | "b_pat (PVar x) = {atom x}"
       
   510 | "b_pat (PUnit) = {}"
       
   511 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
       
   512 | "b_fnclauses (S fc) = (b_fnclause fc)"
       
   513 | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)"
       
   514 | "b_lrb (Clause fcs) = (b_fnclauses fcs)"
       
   515 | "b_fnclause (K x pat exp) = {atom x}"
       
   516 
       
   517 typ exp_raw 
       
   518 thm exp_raw_fnclause_raw_fnclauses_raw_lrb_raw_lrbs_raw_pat_raw.induct[no_vars]
       
   519 thm b_fnclause_raw_b_fnclauses_raw_b_lrb_raw_b_lrbs_raw_b_pat_raw.simps[no_vars]
       
   520 
       
   521 text {*
       
   522   Why does it take so long to define the nominal
       
   523   datatype? Is it the function definitions?
       
   524   
       
   525 *}
       
   526 
       
   527 
       
   528 end
       
   529 
       
   530 
       
   531