Nominal/nominal_dt_rawfuns.ML
changeset 2476 8f8652a8107f
parent 2464 f4eba60cbd69
child 2524 693562f03eee
equal deleted inserted replaced
2475:486d4647bb37 2476:8f8652a8107f
    80 
    80 
    81 
    81 
    82 (* functions for producing sets, fsets and lists of general atom type
    82 (* functions for producing sets, fsets and lists of general atom type
    83    out from concrete atom types *)
    83    out from concrete atom types *)
    84 fun mk_atom_set t =
    84 fun mk_atom_set t =
    85 let
    85   let
    86   val ty = fastype_of t;
    86     val ty = fastype_of t;
    87   val atom_ty = HOLogic.dest_setT ty --> @{typ "atom"};
    87     val atom_ty = HOLogic.dest_setT ty --> @{typ "atom"};
    88   val img_ty = atom_ty --> ty --> @{typ "atom set"};
    88     val img_ty = atom_ty --> ty --> @{typ "atom set"};
    89 in
    89   in
    90   Const (@{const_name image}, img_ty) $ mk_atom_ty atom_ty t
    90     Const (@{const_name image}, img_ty) $ mk_atom_ty atom_ty t
    91 end
    91   end
    92 
    92 
    93 
    93 
    94 fun dest_fsetT (Type (@{type_name fset}, [T])) = T
    94 fun dest_fsetT (Type (@{type_name fset}, [T])) = T
    95   | dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []);
    95   | dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []);
    96 
    96 
    97 fun mk_atom_fset t =
    97 fun mk_atom_fset t =
    98 let
    98   let
    99   val ty = fastype_of t;
    99     val ty = fastype_of t;
   100   val atom_ty = dest_fsetT ty --> @{typ "atom"};
   100     val atom_ty = dest_fsetT ty --> @{typ "atom"};
   101   val fmap_ty = atom_ty --> ty --> @{typ "atom fset"};
   101     val fmap_ty = atom_ty --> ty --> @{typ "atom fset"};
   102   val fset_to_set = @{term "fset_to_set :: atom fset => atom set"}
   102     val fset_to_set = @{term "fset_to_set :: atom fset => atom set"}
   103 in
   103   in
   104   fset_to_set $ (Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t)
   104     fset_to_set $ (Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t)
   105 end
   105   end
   106 
   106 
   107 fun mk_atom_list t =
   107   fun mk_atom_list t =
   108 let
   108   let
   109   val ty = fastype_of t;
   109     val ty = fastype_of t;
   110   val atom_ty = dest_listT ty --> @{typ atom};
   110     val atom_ty = dest_listT ty --> @{typ atom};
   111   val map_ty = atom_ty --> ty --> @{typ "atom list"};
   111     val map_ty = atom_ty --> ty --> @{typ "atom list"};
   112 in
   112   in
   113   Const (@{const_name map}, map_ty) $ mk_atom_ty atom_ty t
   113     Const (@{const_name map}, map_ty) $ mk_atom_ty atom_ty t
   114 end
   114   end
   115 
   115 
   116 (* functions that coerces singletons, sets and fsets of concrete atoms
   116 (* functions that coerces singletons, sets and fsets of concrete atoms
   117    into sets of general atoms *)
   117    into sets of general atoms *)
   118 fun setify ctxt t =
   118 fun setify ctxt t =
   119 let
   119   let
   120   val ty = fastype_of t;
   120     val ty = fastype_of t;
   121 in
   121   in
   122   if is_atom ctxt ty
   122     if is_atom ctxt ty
   123     then  HOLogic.mk_set @{typ atom} [mk_atom t]
   123       then  HOLogic.mk_set @{typ atom} [mk_atom t]
   124   else if is_atom_set ctxt ty
   124     else if is_atom_set ctxt ty
   125     then mk_atom_set t
   125       then mk_atom_set t
   126   else if is_atom_fset ctxt ty
   126     else if is_atom_fset ctxt ty
   127     then mk_atom_fset t
   127       then mk_atom_fset t
   128   else raise TERM ("setify", [t])
   128     else raise TERM ("setify", [t])
   129 end
   129   end
   130 
   130 
   131 (* functions that coerces singletons and lists of concrete atoms
   131 (* functions that coerces singletons and lists of concrete atoms
   132    into lists of general atoms  *)
   132    into lists of general atoms  *)
   133 fun listify ctxt t =
   133 fun listify ctxt t =
   134 let
   134   let
   135   val ty = fastype_of t;
   135     val ty = fastype_of t;
   136 in
   136   in
   137   if is_atom ctxt ty
   137     if is_atom ctxt ty
   138     then HOLogic.mk_list @{typ atom} [mk_atom t]
   138       then HOLogic.mk_list @{typ atom} [mk_atom t]
   139   else if is_atom_list ctxt ty
   139     else if is_atom_list ctxt ty
   140     then mk_atom_set t
   140       then mk_atom_set t
   141   else raise TERM ("listify", [t])
   141     else raise TERM ("listify", [t])
   142 end
   142   end
   143 
   143 
   144 (* coerces a list into a set *)
   144 (* coerces a list into a set *)
   145 fun to_set t =
   145 fun to_set t =
   146   if fastype_of t = @{typ "atom list"}
   146   if fastype_of t = @{typ "atom list"}
   147   then @{term "set::atom list => atom set"} $ t
   147   then @{term "set::atom list => atom set"} $ t
   150 
   150 
   151 
   151 
   152 (** functions that construct the equations for fv and fv_bn **)
   152 (** functions that construct the equations for fv and fv_bn **)
   153 
   153 
   154 fun mk_fv_rhs lthy fv_map fv_bn_map args (BC (bmode, binders, bodies)) =
   154 fun mk_fv_rhs lthy fv_map fv_bn_map args (BC (bmode, binders, bodies)) =
   155 let
   155   let
   156   fun mk_fv_body fv_map args i = 
   156     fun mk_fv_body fv_map args i = 
   157   let
   157       let
   158     val arg = nth args i
   158         val arg = nth args i
   159     val ty = fastype_of arg
   159         val ty = fastype_of arg
   160   in
   160       in
   161     case AList.lookup (op=) fv_map ty of
   161         case AList.lookup (op=) fv_map ty of
   162       NONE => mk_supp arg
   162           NONE => mk_supp arg
   163     | SOME fv => fv $ arg
   163         | SOME fv => fv $ arg
   164   end  
   164       end  
   165 
   165 
   166   fun mk_fv_binder lthy fv_bn_map args binders = 
   166   fun mk_fv_binder lthy fv_bn_map args binders = 
   167   let
   167     let
   168     fun bind_set lthy args (NONE, i) = (setify lthy (nth args i), @{term "{}::atom set"})
   168       fun bind_set lthy args (NONE, i) = (setify lthy (nth args i), @{term "{}::atom set"})
   169       | bind_set _ args (SOME bn, i) = (bn $ (nth args i), 
   169         | bind_set _ args (SOME bn, i) = (bn $ (nth args i), 
   170           if  member (op=) bodies i then @{term "{}::atom set"}  
   170             if  member (op=) bodies i then @{term "{}::atom set"}  
   171           else lookup fv_bn_map bn $ (nth args i))
   171             else lookup fv_bn_map bn $ (nth args i))
   172     fun bind_lst lthy args (NONE, i) = (listify lthy (nth args i), @{term "[]::atom list"})
   172       fun bind_lst lthy args (NONE, i) = (listify lthy (nth args i), @{term "[]::atom list"})
   173       | bind_lst _ args (SOME bn, i) = (bn $ (nth args i),
   173         | bind_lst _ args (SOME bn, i) = (bn $ (nth args i),
   174           if  member (op=) bodies i then @{term "[]::atom list"}  
   174             if  member (op=) bodies i then @{term "[]::atom list"}  
   175           else lookup fv_bn_map bn $ (nth args i)) 
   175             else lookup fv_bn_map bn $ (nth args i)) 
   176   
   176   
   177     val (combine_fn, bind_fn) =
   177       val (combine_fn, bind_fn) =
   178       case bmode of
   178         case bmode of
   179         Lst => (fold_append, bind_lst) 
   179           Lst => (fold_append, bind_lst) 
   180       | Set => (fold_union, bind_set)
   180         | Set => (fold_union, bind_set)
   181       | Res => (fold_union, bind_set)
   181         | Res => (fold_union, bind_set)
   182   in
   182     in
   183     binders
   183       binders
   184     |> map (bind_fn lthy args)
   184       |> map (bind_fn lthy args)
   185     |> split_list
   185       |> split_list
   186     |> pairself combine_fn
   186       |> pairself combine_fn
   187   end  
   187     end  
   188 
   188 
   189   val t1 = map (mk_fv_body fv_map args) bodies
   189     val t1 = map (mk_fv_body fv_map args) bodies
   190   val (t2, t3) = mk_fv_binder lthy fv_bn_map args binders
   190     val (t2, t3) = mk_fv_binder lthy fv_bn_map args binders
   191 in 
   191   in 
   192   mk_union (mk_diff (fold_union t1, to_set t2), to_set t3)
   192     mk_union (mk_diff (fold_union t1, to_set t2), to_set t3)
   193 end
   193   end
   194 
   194 
   195 (* in case of fv_bn we have to treat the case special, where an
   195 (* in case of fv_bn we have to treat the case special, where an
   196    "empty" binding clause is given *)
   196    "empty" binding clause is given *)
   197 fun mk_fv_bn_rhs lthy fv_map fv_bn_map bn_args args bclause =
   197 fun mk_fv_bn_rhs lthy fv_map fv_bn_map bn_args args bclause =
   198 let
   198   let
   199   fun mk_fv_bn_body fv_map fv_bn_map bn_args args i = 
   199     fun mk_fv_bn_body i = 
   200   let
   200     let
   201     val arg = nth args i
   201       val arg = nth args i
   202     val ty = fastype_of arg
   202       val ty = fastype_of arg
   203   in
   203     in
   204     case AList.lookup (op=) bn_args i of
   204       case AList.lookup (op=) bn_args i of
   205       NONE => (case (AList.lookup (op=) fv_map ty) of
   205         NONE => (case (AList.lookup (op=) fv_map ty) of
   206                  NONE => mk_supp arg
   206                    NONE => mk_supp arg
   207                | SOME fv => fv $ arg)
   207                  | SOME fv => fv $ arg)
   208     | SOME (NONE) => @{term "{}::atom set"}
   208       | SOME (NONE) => @{term "{}::atom set"}
   209     | SOME (SOME bn) => lookup fv_bn_map bn $ arg
   209       | SOME (SOME bn) => lookup fv_bn_map bn $ arg
   210   end  
   210     end  
   211 in
   211   in
   212   case bclause of
   212     case bclause of
   213     BC (_, [], bodies) => fold_union (map (mk_fv_bn_body fv_map fv_bn_map bn_args args) bodies)
   213       BC (_, [], bodies) => fold_union (map mk_fv_bn_body bodies)
   214   | _ => mk_fv_rhs lthy fv_map fv_bn_map args bclause
   214     | _ => mk_fv_rhs lthy fv_map fv_bn_map args bclause
   215 end
   215   end
   216 
   216 
   217 fun mk_fv_eq lthy fv_map fv_bn_map (constr, ty, arg_tys, _) bclauses = 
   217 fun mk_fv_eq lthy fv_map fv_bn_map (constr, ty, arg_tys, _) bclauses = 
   218 let
   218   let
   219   val arg_names = Datatype_Prop.make_tnames arg_tys
   219     val arg_names = Datatype_Prop.make_tnames arg_tys
   220   val args = map Free (arg_names ~~ arg_tys)
   220     val args = map Free (arg_names ~~ arg_tys)
   221   val fv = lookup fv_map ty
   221     val fv = lookup fv_map ty
   222   val lhs = fv $ list_comb (constr, args)
   222     val lhs = fv $ list_comb (constr, args)
   223   val rhs_trms = map (mk_fv_rhs lthy fv_map fv_bn_map args) bclauses
   223     val rhs_trms = map (mk_fv_rhs lthy fv_map fv_bn_map args) bclauses
   224   val rhs = fold_union rhs_trms
   224     val rhs = fold_union rhs_trms
   225 in
   225   in
   226   HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
   226     HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
   227 end
   227   end
   228 
   228 
   229 fun mk_fv_bn_eq lthy bn_trm fv_map fv_bn_map (bn_args, (constr, _, arg_tys, _)) bclauses =
   229 fun mk_fv_bn_eq lthy bn_trm fv_map fv_bn_map (bn_args, (constr, _, arg_tys, _)) bclauses =
   230 let
   230   let
   231   val arg_names = Datatype_Prop.make_tnames arg_tys
   231     val arg_names = Datatype_Prop.make_tnames arg_tys
   232   val args = map Free (arg_names ~~ arg_tys)
   232     val args = map Free (arg_names ~~ arg_tys)
   233   val fv_bn = lookup fv_bn_map bn_trm
   233     val fv_bn = lookup fv_bn_map bn_trm
   234   val lhs = fv_bn $ list_comb (constr, args)
   234     val lhs = fv_bn $ list_comb (constr, args)
   235   val rhs_trms = map (mk_fv_bn_rhs lthy fv_map fv_bn_map bn_args args) bclauses
   235     val rhs_trms = map (mk_fv_bn_rhs lthy fv_map fv_bn_map bn_args args) bclauses
   236   val rhs = fold_union rhs_trms
   236     val rhs = fold_union rhs_trms
   237 in
   237   in
   238   HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
   238     HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
   239 end
   239   end
   240 
   240 
   241 fun mk_fv_bn_eqs lthy fv_map fv_bn_map constrs_info bclausesss (bn_trm, bn_n, bn_argss) = 
   241 fun mk_fv_bn_eqs lthy fv_map fv_bn_map constrs_info bclausesss (bn_trm, bn_n, bn_argss) = 
   242 let
   242   let
   243   val nth_constrs_info = nth constrs_info bn_n
   243     val nth_constrs_info = nth constrs_info bn_n
   244   val nth_bclausess = nth bclausesss bn_n
   244     val nth_bclausess = nth bclausesss bn_n
   245 in
   245   in
   246   map2 (mk_fv_bn_eq lthy bn_trm fv_map fv_bn_map) (bn_argss ~~ nth_constrs_info) nth_bclausess
   246     map2 (mk_fv_bn_eq lthy bn_trm fv_map fv_bn_map) (bn_argss ~~ nth_constrs_info) nth_bclausess
   247 end
   247   end
   248 
   248 
   249 fun define_raw_fvs raw_full_ty_names raw_tys cns_info bn_info bclausesss constr_thms size_simps lthy =
   249 fun define_raw_fvs raw_full_ty_names raw_tys cns_info bn_info bclausesss constr_thms size_simps lthy =
   250 let
   250   let
   251   val fv_names = map (prefix "fv_" o Long_Name.base_name) raw_full_ty_names
   251     val fv_names = map (prefix "fv_" o Long_Name.base_name) raw_full_ty_names
   252   val fv_tys = map (fn ty => ty --> @{typ "atom set"}) raw_tys
   252     val fv_tys = map (fn ty => ty --> @{typ "atom set"}) raw_tys
   253   val fv_frees = map Free (fv_names ~~ fv_tys);
   253     val fv_frees = map Free (fv_names ~~ fv_tys);
   254   val fv_map = raw_tys ~~ fv_frees
   254     val fv_map = raw_tys ~~ fv_frees
   255 
   255 
   256   val (bns, bn_tys) = split_list (map (fn (bn, i, _) => (bn, i)) bn_info)
   256     val (bns, bn_tys) = split_list (map (fn (bn, i, _) => (bn, i)) bn_info)
   257   val bn_names = map (fn bn => Long_Name.base_name (fst (dest_Const bn))) bns
   257     val bn_names = map (fn bn => Long_Name.base_name (fst (dest_Const bn))) bns
   258   val fv_bn_names = map (prefix "fv_") bn_names
   258     val fv_bn_names = map (prefix "fv_") bn_names
   259   val fv_bn_arg_tys = map (nth raw_tys) bn_tys
   259     val fv_bn_arg_tys = map (nth raw_tys) bn_tys
   260   val fv_bn_tys = map (fn ty => ty --> @{typ "atom set"}) fv_bn_arg_tys
   260     val fv_bn_tys = map (fn ty => ty --> @{typ "atom set"}) fv_bn_arg_tys
   261   val fv_bn_frees = map Free (fv_bn_names ~~ fv_bn_tys)
   261     val fv_bn_frees = map Free (fv_bn_names ~~ fv_bn_tys)
   262   val fv_bn_map = bns ~~ fv_bn_frees
   262     val fv_bn_map = bns ~~ fv_bn_frees
   263 
   263 
   264   val fv_eqs = map2 (map2 (mk_fv_eq lthy fv_map fv_bn_map)) cns_info bclausesss 
   264     val fv_eqs = map2 (map2 (mk_fv_eq lthy fv_map fv_bn_map)) cns_info bclausesss 
   265   val fv_bn_eqs = map (mk_fv_bn_eqs lthy fv_map fv_bn_map cns_info bclausesss) bn_info
   265     val fv_bn_eqs = map (mk_fv_bn_eqs lthy fv_map fv_bn_map cns_info bclausesss) bn_info
   266   
   266   
   267   val all_fun_names = map (fn s => (Binding.name s, NONE, NoSyn)) (fv_names @ fv_bn_names)
   267     val all_fun_names = map (fn s => (Binding.name s, NONE, NoSyn)) (fv_names @ fv_bn_names)
   268   val all_fun_eqs = map (pair Attrib.empty_binding) (flat fv_eqs @ flat fv_bn_eqs)
   268     val all_fun_eqs = map (pair Attrib.empty_binding) (flat fv_eqs @ flat fv_bn_eqs)
   269 
   269 
   270   val (_, lthy') = Function.add_function all_fun_names all_fun_eqs
   270     val (_, lthy') = Function.add_function all_fun_names all_fun_eqs
   271     Function_Common.default_config (pat_completeness_simp constr_thms) lthy
   271       Function_Common.default_config (pat_completeness_simp constr_thms) lthy
   272   
   272   
   273   val (info, lthy'') = prove_termination size_simps (Local_Theory.restore lthy')
   273     val (info, lthy'') = prove_termination size_simps (Local_Theory.restore lthy')
   274  
   274  
   275   val {fs, simps, inducts, ...} = info;
   275     val {fs, simps, inducts, ...} = info;
   276 
   276 
   277   val morphism = ProofContext.export_morphism lthy'' lthy
   277     val morphism = ProofContext.export_morphism lthy'' lthy
   278   val simps_exp = map (Morphism.thm morphism) (the simps)
   278     val simps_exp = map (Morphism.thm morphism) (the simps)
   279   val inducts_exp = map (Morphism.thm morphism) (the inducts)
   279     val inducts_exp = map (Morphism.thm morphism) (the inducts)
   280 
   280 
   281   val (fvs', fv_bns') = chop (length fv_frees) fs
   281     val (fvs', fv_bns') = chop (length fv_frees) fs
   282 in
   282   in
   283   (fvs', fv_bns', simps_exp, inducts_exp, lthy'')
   283     (fvs', fv_bns', simps_exp, inducts_exp, lthy'')
   284 end
   284   end
   285 
   285 
   286 
   286 
   287 (** equivarance proofs **)
   287 (** equivarance proofs **)
   288 
   288 
   289 val eqvt_apply_sym = @{thm eqvt_apply[symmetric]}
   289 val eqvt_apply_sym = @{thm eqvt_apply[symmetric]}
   300     (Object_Logic.full_atomize_tac
   300     (Object_Logic.full_atomize_tac
   301      THEN' (DETERM o (InductTacs.induct_rules_tac ctxt insts ind_thms))  
   301      THEN' (DETERM o (InductTacs.induct_rules_tac ctxt insts ind_thms))  
   302      THEN_ALL_NEW  subproof_tac const_names simps ctxt)
   302      THEN_ALL_NEW  subproof_tac const_names simps ctxt)
   303 
   303 
   304 fun mk_eqvt_goal pi const arg =
   304 fun mk_eqvt_goal pi const arg =
   305 let
   305   let
   306   val lhs = mk_perm pi (const $ arg)
   306     val lhs = mk_perm pi (const $ arg)
   307   val rhs = const $ (mk_perm pi arg)  
   307     val rhs = const $ (mk_perm pi arg)  
   308 in
   308   in
   309   HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
   309     HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
   310 end
   310   end
   311 
   311 
   312 fun raw_prove_eqvt consts ind_thms simps ctxt =
   312 fun raw_prove_eqvt consts ind_thms simps ctxt =
   313   if null consts then []
   313   if null consts then []
   314   else
   314   else
   315     let 
   315     let