Nominal/Nominal2.thy
changeset 2647 5e95387bef45
parent 2643 0579d3a48304
child 2649 a8ebcb368a15
equal deleted inserted replaced
2646:51f75d24bd73 2647:5e95387bef45
   161 in
   161 in
   162   (raw_dt_full_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, cnstr_names, lthy1)
   162   (raw_dt_full_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, cnstr_names, lthy1)
   163 end
   163 end
   164 *}
   164 *}
   165 
   165 
   166 ML {*
       
   167 (* for testing porposes - to exit the procedure early *)
       
   168 exception TEST of Proof.context
       
   169 
       
   170 val (STEPS, STEPS_setup) = Attrib.config_int "STEPS" (K 100);
       
   171 
       
   172 fun get_STEPS ctxt = Config.get ctxt STEPS
       
   173 *}
       
   174 
       
   175 
       
   176 setup STEPS_setup
       
   177 
   166 
   178 ML {*
   167 ML {*
   179 fun nominal_datatype2 opt_thms_name dts bn_funs bn_eqs bclauses lthy =
   168 fun nominal_datatype2 opt_thms_name dts bn_funs bn_eqs bclauses lthy =
   180 let
   169 let
   181   (* definition of the raw datatypes *)
   170   (* definition of the raw datatypes *)
   182   val _ = warning "Definition of raw datatypes";
   171   val _ = trace_msg (K "Defining raw datatypes...")
   183   val (raw_dt_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, cnstr_names, lthy0) =
   172   val (raw_dt_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, cnstr_names, lthy0) =
   184     if get_STEPS lthy > 0 
   173     define_raw_dts dts bn_funs bn_eqs bclauses lthy   
   185     then define_raw_dts dts bn_funs bn_eqs bclauses lthy
       
   186     else raise TEST lthy
       
   187 
   174 
   188   val dtinfo = Datatype.the_info (ProofContext.theory_of lthy0) (hd raw_dt_names)
   175   val dtinfo = Datatype.the_info (ProofContext.theory_of lthy0) (hd raw_dt_names)
   189   val {descr, sorts, ...} = dtinfo
   176   val {descr, sorts, ...} = dtinfo
   190 
   177 
   191   val raw_tys = all_dtyps descr sorts
   178   val raw_tys = all_dtyps descr sorts
   207   val raw_size_trms = map HOLogic.size_const raw_tys
   194   val raw_size_trms = map HOLogic.size_const raw_tys
   208   val raw_size_thms = Size.size_thms (ProofContext.theory_of lthy0) (hd raw_dt_names)
   195   val raw_size_thms = Size.size_thms (ProofContext.theory_of lthy0) (hd raw_dt_names)
   209     |> `(fn thms => (length thms) div 2)
   196     |> `(fn thms => (length thms) div 2)
   210     |> uncurry drop
   197     |> uncurry drop
   211   
   198   
   212   (* definitions of raw permutations by primitive recursion *)
   199   val _ = trace_msg (K "Defining raw permutations...")
   213   val _ = warning "Definition of raw permutations";
       
   214   val ((raw_perm_funs, raw_perm_simps, raw_perm_laws), lthy2a) =
   200   val ((raw_perm_funs, raw_perm_simps, raw_perm_laws), lthy2a) =
   215     if get_STEPS lthy0 > 0 
   201     define_raw_perms raw_full_ty_names raw_tys tvs (flat raw_constrs) raw_induct_thm lthy0
   216     then define_raw_perms raw_full_ty_names raw_tys tvs (flat raw_constrs) raw_induct_thm lthy0
       
   217     else raise TEST lthy0
       
   218  
   202  
   219   (* noting the raw permutations as eqvt theorems *)
   203   (* noting the raw permutations as eqvt theorems *)
   220   val (_, lthy3) = Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_perm_simps) lthy2a
   204   val (_, lthy3) = Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_perm_simps) lthy2a
   221 
   205 
   222   (* definition of raw fv and bn functions *)
   206 
   223   val _ = warning "Definition of raw fv- and bn-functions";
   207   val _ = trace_msg (K "Defining raw fv- and bn-functions...")
   224   val (raw_bns, raw_bn_defs, raw_bn_info, raw_bn_induct, lthy3a) =
   208   val (raw_bns, raw_bn_defs, raw_bn_info, raw_bn_induct, lthy3a) =
   225     if get_STEPS lthy3 > 1
   209     define_raw_bns raw_full_ty_names raw_dts raw_bn_funs raw_bn_eqs 
   226     then define_raw_bns raw_full_ty_names raw_dts raw_bn_funs raw_bn_eqs 
       
   227       (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3
   210       (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3
   228     else raise TEST lthy3
   211     
   229 
       
   230   (* defining the permute_bn functions *)
   212   (* defining the permute_bn functions *)
   231   val (raw_perm_bns, raw_perm_bn_simps, lthy3b) = 
   213   val (raw_perm_bns, raw_perm_bn_simps, lthy3b) = 
   232     if get_STEPS lthy3a > 2
   214     define_raw_bn_perms raw_tys raw_bn_info raw_cns_info 
   233     then define_raw_bn_perms raw_tys raw_bn_info raw_cns_info 
       
   234       (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3a
   215       (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3a
   235     else raise TEST lthy3a
   216     
   236 
       
   237   val (raw_fvs, raw_fv_bns, raw_fv_defs, raw_fv_bns_induct, lthy3c) = 
   217   val (raw_fvs, raw_fv_bns, raw_fv_defs, raw_fv_bns_induct, lthy3c) = 
   238     if get_STEPS lthy3b > 3 
   218     define_raw_fvs raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses 
   239     then define_raw_fvs raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses 
       
   240       (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3b
   219       (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3b
   241     else raise TEST lthy3b
   220     
   242 
   221   val _ = trace_msg (K "Defining alpha relations...")
   243   (* definition of raw alphas *)
       
   244   val _ = warning "Definition of alphas";
       
   245   val (alpha_trms, alpha_bn_trms, alpha_intros, alpha_cases, alpha_induct, lthy4) =
   222   val (alpha_trms, alpha_bn_trms, alpha_intros, alpha_cases, alpha_induct, lthy4) =
   246     if get_STEPS lthy3c > 4 
   223     define_raw_alpha raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses raw_fvs lthy3c
   247     then define_raw_alpha raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses raw_fvs lthy3c
   224     
   248     else raise TEST lthy3c
       
   249   val alpha_tys = map (domain_type o fastype_of) alpha_trms  
   225   val alpha_tys = map (domain_type o fastype_of) alpha_trms  
   250 
   226 
   251   (* definition of alpha-distinct lemmas *)
   227   val _ = trace_msg (K "Proving distinct theorems...")
   252   val _ = warning "Distinct theorems";
       
   253   val alpha_distincts = 
   228   val alpha_distincts = 
   254     mk_alpha_distincts lthy4 alpha_cases raw_distinct_thms alpha_trms raw_tys
   229     mk_alpha_distincts lthy4 alpha_cases raw_distinct_thms alpha_trms raw_tys
   255 
   230 
   256   (* definition of alpha_eq_iff  lemmas *)
   231   val _ = trace_msg (K "Proving eq-iff theorems...")
   257   val _ = warning "Eq-iff theorems";
       
   258   val alpha_eq_iff = 
   232   val alpha_eq_iff = 
   259     if get_STEPS lthy > 5
   233     mk_alpha_eq_iff lthy4 alpha_intros raw_distinct_thms raw_inject_thms alpha_cases
   260     then mk_alpha_eq_iff lthy4 alpha_intros raw_distinct_thms raw_inject_thms alpha_cases
   234     
   261     else raise TEST lthy4
   235   val _ = trace_msg (K "Proving equivariance of bns, fvs, size and alpha...")
   262 
       
   263   (* proving equivariance lemmas for bns, fvs, size and alpha *)
       
   264   val _ = warning "Proving equivariance";
       
   265   val raw_bn_eqvt = 
   236   val raw_bn_eqvt = 
   266     if get_STEPS lthy > 6
   237     raw_prove_eqvt raw_bns raw_bn_induct (raw_bn_defs @ raw_perm_simps) lthy4
   267     then raw_prove_eqvt raw_bns raw_bn_induct (raw_bn_defs @ raw_perm_simps) lthy4
   238     
   268     else raise TEST lthy4
       
   269 
       
   270   (* noting the raw_bn_eqvt lemmas in a temprorary theory *)
   239   (* noting the raw_bn_eqvt lemmas in a temprorary theory *)
   271   val lthy_tmp = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_bn_eqvt) lthy4)
   240   val lthy_tmp = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_bn_eqvt) lthy4)
   272 
   241 
   273   val raw_fv_eqvt = 
   242   val raw_fv_eqvt = 
   274     if get_STEPS lthy > 7
   243     raw_prove_eqvt (raw_fvs @ raw_fv_bns) raw_fv_bns_induct (raw_fv_defs @ raw_perm_simps) 
   275     then raw_prove_eqvt (raw_fvs @ raw_fv_bns) raw_fv_bns_induct (raw_fv_defs @ raw_perm_simps) 
       
   276       (Local_Theory.restore lthy_tmp)
   244       (Local_Theory.restore lthy_tmp)
   277     else raise TEST lthy4
   245     
   278 
       
   279   val raw_size_eqvt = 
   246   val raw_size_eqvt = 
   280     if get_STEPS lthy > 8
   247     raw_prove_eqvt raw_size_trms raw_induct_thms (raw_size_thms @ raw_perm_simps) 
   281     then raw_prove_eqvt raw_size_trms raw_induct_thms (raw_size_thms @ raw_perm_simps) 
       
   282       (Local_Theory.restore lthy_tmp)
   248       (Local_Theory.restore lthy_tmp)
   283       |> map (rewrite_rule @{thms permute_nat_def[THEN eq_reflection]})
   249       |> map (rewrite_rule @{thms permute_nat_def[THEN eq_reflection]})
   284       |> map (fn thm => thm RS @{thm sym})
   250       |> map (fn thm => thm RS @{thm sym})
   285     else raise TEST lthy4
   251      
   286  
       
   287   val lthy5 = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_fv_eqvt) lthy_tmp)
   252   val lthy5 = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_fv_eqvt) lthy_tmp)
   288 
   253 
   289   val (alpha_eqvt, lthy6) =
   254   val (alpha_eqvt, lthy6) =
   290     if get_STEPS lthy > 9
   255     Nominal_Eqvt.equivariance true (alpha_trms @ alpha_bn_trms) alpha_induct alpha_intros lthy5
   291     then Nominal_Eqvt.equivariance true (alpha_trms @ alpha_bn_trms) alpha_induct alpha_intros lthy5
   256 
   292     else raise TEST lthy4
   257   val _ = trace_msg (K "Proving equivalence of alpha...")
   293 
       
   294   (* proving alpha equivalence *)
       
   295   val _ = warning "Proving equivalence"
       
   296 
       
   297   val alpha_refl_thms = 
   258   val alpha_refl_thms = 
   298     if get_STEPS lthy > 10
   259     raw_prove_refl alpha_trms alpha_bn_trms alpha_intros raw_induct_thm lthy6 
   299     then raw_prove_refl alpha_trms alpha_bn_trms alpha_intros raw_induct_thm lthy6 
       
   300     else raise TEST lthy6
       
   301 
   260 
   302   val alpha_sym_thms = 
   261   val alpha_sym_thms = 
   303     if get_STEPS lthy > 11
   262     raw_prove_sym (alpha_trms @ alpha_bn_trms) alpha_intros alpha_induct lthy6 
   304     then raw_prove_sym (alpha_trms @ alpha_bn_trms) alpha_intros alpha_induct lthy6 
       
   305     else raise TEST lthy6
       
   306 
   263 
   307   val alpha_trans_thms = 
   264   val alpha_trans_thms = 
   308     if get_STEPS lthy > 12
   265     raw_prove_trans (alpha_trms @ alpha_bn_trms) (raw_distinct_thms @ raw_inject_thms) 
   309     then raw_prove_trans (alpha_trms @ alpha_bn_trms) (raw_distinct_thms @ raw_inject_thms) 
   266       alpha_intros alpha_induct alpha_cases lthy6
   310            alpha_intros alpha_induct alpha_cases lthy6
       
   311     else raise TEST lthy6
       
   312 
   267 
   313   val (alpha_equivp_thms, alpha_bn_equivp_thms) = 
   268   val (alpha_equivp_thms, alpha_bn_equivp_thms) = 
   314     if get_STEPS lthy > 13
   269     raw_prove_equivp alpha_trms alpha_bn_trms alpha_refl_thms alpha_sym_thms 
   315     then raw_prove_equivp alpha_trms alpha_bn_trms alpha_refl_thms alpha_sym_thms 
   270       alpha_trans_thms lthy6
   316        alpha_trans_thms lthy6
   271 
   317     else raise TEST lthy6
   272   val _ = trace_msg (K "Proving alpha implies bn...")
   318 
       
   319   (* proving alpha implies alpha_bn *)
       
   320   val _ = warning "Proving alpha implies bn"
       
   321 
       
   322   val alpha_bn_imp_thms = 
   273   val alpha_bn_imp_thms = 
   323     if get_STEPS lthy > 14
   274     raw_prove_bn_imp alpha_trms alpha_bn_trms alpha_intros alpha_induct lthy6 
   324     then raw_prove_bn_imp alpha_trms alpha_bn_trms alpha_intros alpha_induct lthy6 
   275   
   325     else raise TEST lthy6
   276   val _ = trace_msg (K "Proving respectfulness...")
   326   
       
   327   (* respectfulness proofs *)
       
   328   val raw_funs_rsp_aux = 
   277   val raw_funs_rsp_aux = 
   329     if get_STEPS lthy > 15
   278     raw_fv_bn_rsp_aux alpha_trms alpha_bn_trms raw_fvs 
   330     then raw_fv_bn_rsp_aux alpha_trms alpha_bn_trms raw_fvs 
       
   331       raw_bns raw_fv_bns alpha_induct (raw_bn_defs @ raw_fv_defs) lthy6
   279       raw_bns raw_fv_bns alpha_induct (raw_bn_defs @ raw_fv_defs) lthy6
   332     else raise TEST lthy6 
   280   
   333   
   281   val raw_funs_rsp = map mk_funs_rsp raw_funs_rsp_aux
   334   val raw_funs_rsp = 
       
   335     if get_STEPS lthy > 16
       
   336     then map mk_funs_rsp raw_funs_rsp_aux
       
   337     else raise TEST lthy6 
       
   338 
   282 
   339   val raw_size_rsp = 
   283   val raw_size_rsp = 
   340     if get_STEPS lthy > 17
   284     raw_size_rsp_aux (alpha_trms @ alpha_bn_trms) alpha_induct 
   341     then
       
   342       raw_size_rsp_aux (alpha_trms @ alpha_bn_trms) alpha_induct 
       
   343       (raw_size_thms @ raw_size_eqvt) lthy6
   285       (raw_size_thms @ raw_size_eqvt) lthy6
   344       |> map mk_funs_rsp
   286       |> map mk_funs_rsp
   345     else raise TEST lthy6 
       
   346 
   287 
   347   val raw_constrs_rsp = 
   288   val raw_constrs_rsp = 
   348     if get_STEPS lthy > 18
   289     raw_constrs_rsp (flat raw_constrs) alpha_trms alpha_intros
   349     then raw_constrs_rsp (flat raw_constrs) alpha_trms alpha_intros
       
   350       (alpha_bn_imp_thms @ raw_funs_rsp_aux) lthy6 
   290       (alpha_bn_imp_thms @ raw_funs_rsp_aux) lthy6 
   351     else raise TEST lthy6 
   291     
   352 
   292   val alpha_permute_rsp = map mk_alpha_permute_rsp alpha_eqvt
   353   val alpha_permute_rsp = 
       
   354     if get_STEPS lthy > 19
       
   355     then map mk_alpha_permute_rsp alpha_eqvt
       
   356     else raise TEST lthy6 
       
   357 
   293 
   358   val alpha_bn_rsp = 
   294   val alpha_bn_rsp = 
   359     if get_STEPS lthy > 20
   295     raw_alpha_bn_rsp alpha_bn_trms alpha_bn_equivp_thms alpha_bn_imp_thms
   360     then raw_alpha_bn_rsp alpha_bn_trms alpha_bn_equivp_thms alpha_bn_imp_thms
       
   361     else raise TEST lthy6 
       
   362 
   296 
   363   val raw_perm_bn_rsp =
   297   val raw_perm_bn_rsp =
   364     if get_STEPS lthy > 21
   298     raw_perm_bn_rsp (alpha_trms @ alpha_bn_trms) raw_perm_bns alpha_induct 
   365     then raw_perm_bn_rsp (alpha_trms @ alpha_bn_trms) raw_perm_bns alpha_induct 
       
   366       alpha_intros raw_perm_bn_simps lthy6
   299       alpha_intros raw_perm_bn_simps lthy6
   367     else raise TEST lthy6
       
   368 
   300 
   369   (* noting the quot_respects lemmas *)
   301   (* noting the quot_respects lemmas *)
   370   val (_, lthy6a) = 
   302   val (_, lthy6a) = 
   371     if get_STEPS lthy > 22
   303     Local_Theory.note ((Binding.empty, [rsp_attr]),
   372     then Local_Theory.note ((Binding.empty, [rsp_attr]),
       
   373       raw_constrs_rsp @ raw_funs_rsp @ raw_size_rsp @ alpha_permute_rsp @ 
   304       raw_constrs_rsp @ raw_funs_rsp @ raw_size_rsp @ alpha_permute_rsp @ 
   374       alpha_bn_rsp @ raw_perm_bn_rsp) lthy6
   305       alpha_bn_rsp @ raw_perm_bn_rsp) lthy6
   375     else raise TEST lthy6
   306 
   376 
   307   val _ = trace_msg (K "Defining the quotient types...")
   377   (* defining the quotient type *)
       
   378   val _ = warning "Declaring the quotient types"
       
   379   val qty_descr = map (fn (vs, bind, mx, _) => (vs, bind, mx)) dts
   308   val qty_descr = map (fn (vs, bind, mx, _) => (vs, bind, mx)) dts
   380      
   309      
   381   val (qty_infos, lthy7) = 
   310   val (qty_infos, lthy7) = 
   382     if get_STEPS lthy > 23
   311     define_qtypes qty_descr alpha_tys alpha_trms alpha_equivp_thms lthy6a
   383     then define_qtypes qty_descr alpha_tys alpha_trms alpha_equivp_thms lthy6a
       
   384     else raise TEST lthy6a
       
   385 
   312 
   386   val qtys = map #qtyp qty_infos
   313   val qtys = map #qtyp qty_infos
   387   val qty_full_names = map (fst o dest_Type) qtys
   314   val qty_full_names = map (fst o dest_Type) qtys
   388   val qty_names = map Long_Name.base_name qty_full_names             
   315   val qty_names = map Long_Name.base_name qty_full_names             
   389 
   316 
   390   (* defining of quotient term-constructors, binding functions, free vars functions *)
   317   val _ = trace_msg (K "Defining the quotient constants...")
   391   val _ = warning "Defining the quotient constants"
       
   392   val qconstrs_descrs =
   318   val qconstrs_descrs =
   393     (map2 o map2) (fn (b, _, mx) => fn t => (Name.of_binding b, t, mx)) (get_cnstrs dts) raw_constrs
   319     (map2 o map2) (fn (b, _, mx) => fn t => (Name.of_binding b, t, mx)) (get_cnstrs dts) raw_constrs
   394 
   320 
   395   val qbns_descr =
   321   val qbns_descr =
   396     map2 (fn (b, _, mx) => fn t => (Name.of_binding b, t, mx)) bn_funs raw_bns
   322     map2 (fn (b, _, mx) => fn t => (Name.of_binding b, t, mx)) bn_funs raw_bns
   413   val qperm_bn_descr = 
   339   val qperm_bn_descr = 
   414     map2 (fn (b, _, _) => fn t => ("permute_" ^ Name.of_binding b, t, NoSyn)) bn_funs raw_perm_bns
   340     map2 (fn (b, _, _) => fn t => ("permute_" ^ Name.of_binding b, t, NoSyn)) bn_funs raw_perm_bns
   415      
   341      
   416   val ((((((qconstrs_infos, qbns_info), qfvs_info), qfv_bns_info), qalpha_bns_info), qperm_bns_info), 
   342   val ((((((qconstrs_infos, qbns_info), qfvs_info), qfv_bns_info), qalpha_bns_info), qperm_bns_info), 
   417     lthy8) = 
   343     lthy8) = 
   418     if get_STEPS lthy > 24
       
   419     then 
       
   420       lthy7
   344       lthy7
   421       |> fold_map (define_qconsts qtys) qconstrs_descrs 
   345       |> fold_map (define_qconsts qtys) qconstrs_descrs 
   422       ||>> define_qconsts qtys qbns_descr 
   346       ||>> define_qconsts qtys qbns_descr 
   423       ||>> define_qconsts qtys qfvs_descr
   347       ||>> define_qconsts qtys qfvs_descr
   424       ||>> define_qconsts qtys qfv_bns_descr
   348       ||>> define_qconsts qtys qfv_bns_descr
   425       ||>> define_qconsts qtys qalpha_bns_descr
   349       ||>> define_qconsts qtys qalpha_bns_descr
   426       ||>> define_qconsts qtys qperm_bn_descr
   350       ||>> define_qconsts qtys qperm_bn_descr
   427     else raise TEST lthy7
   351 
   428 
       
   429   (* definition of the quotient permfunctions and pt-class *)
       
   430   val lthy9 = 
   352   val lthy9 = 
   431     if get_STEPS lthy > 25
   353     define_qperms qtys qty_full_names tvs qperm_descr raw_perm_laws lthy8 
   432     then define_qperms qtys qty_full_names tvs qperm_descr raw_perm_laws lthy8 
       
   433     else raise TEST lthy8
       
   434   
   354   
   435   val lthy9a = 
   355   val lthy9a = 
   436     if get_STEPS lthy > 26
   356     define_qsizes qtys qty_full_names tvs qsize_descr lthy9
   437     then define_qsizes qtys qty_full_names tvs qsize_descr lthy9
       
   438     else raise TEST lthy9
       
   439 
   357 
   440   val qtrms = (map o map) #qconst qconstrs_infos
   358   val qtrms = (map o map) #qconst qconstrs_infos
   441   val qbns = map #qconst qbns_info
   359   val qbns = map #qconst qbns_info
   442   val qfvs = map #qconst qfvs_info
   360   val qfvs = map #qconst qfvs_info
   443   val qfv_bns = map #qconst qfv_bns_info
   361   val qfv_bns = map #qconst qfv_bns_info
   444   val qalpha_bns = map #qconst qalpha_bns_info
   362   val qalpha_bns = map #qconst qalpha_bns_info
   445   val qperm_bns = map #qconst qperm_bns_info
   363   val qperm_bns = map #qconst qperm_bns_info
   446 
   364 
   447   (* lifting of the theorems *)
   365   val _ = trace_msg (K "Lifting of theorems...")
   448   val _ = warning "Lifting of Theorems"
       
   449   
   366   
   450   val eq_iff_simps = @{thms alphas permute_prod.simps prod_fv.simps prod_alpha_def prod_rel_def
   367   val eq_iff_simps = @{thms alphas permute_prod.simps prod_fv.simps prod_alpha_def prod_rel_def
   451     prod.cases} 
   368     prod.cases} 
   452 
   369 
   453   val ((((((qdistincts, qeq_iffs), qfv_defs), qbn_defs), qperm_simps), qfv_qbn_eqvts), lthyA) = 
   370   val ((((((qdistincts, qeq_iffs), qfv_defs), qbn_defs), qperm_simps), qfv_qbn_eqvts), lthyA) = 
   454     if get_STEPS lthy > 27
   371     lthy9a    
   455     then 
   372     |> lift_thms qtys [] alpha_distincts  
   456       lthy9a    
   373     ||>> lift_thms qtys eq_iff_simps alpha_eq_iff       
   457       |> lift_thms qtys [] alpha_distincts  
   374     ||>> lift_thms qtys [] raw_fv_defs
   458       ||>> lift_thms qtys eq_iff_simps alpha_eq_iff       
   375     ||>> lift_thms qtys [] raw_bn_defs
   459       ||>> lift_thms qtys [] raw_fv_defs
   376     ||>> lift_thms qtys [] raw_perm_simps
   460       ||>> lift_thms qtys [] raw_bn_defs
   377     ||>> lift_thms qtys [] (raw_fv_eqvt @ raw_bn_eqvt)
   461       ||>> lift_thms qtys [] raw_perm_simps
       
   462       ||>> lift_thms qtys [] (raw_fv_eqvt @ raw_bn_eqvt)
       
   463     else raise TEST lthy9a
       
   464 
   378 
   465   val ((((((qsize_eqvt, [qinduct]), qexhausts), qsize_simps), qperm_bn_simps), qalpha_refl_thms), lthyB) = 
   379   val ((((((qsize_eqvt, [qinduct]), qexhausts), qsize_simps), qperm_bn_simps), qalpha_refl_thms), lthyB) = 
   466     if get_STEPS lthy > 28
   380     lthyA 
   467     then
   381     |> lift_thms qtys [] raw_size_eqvt
   468       lthyA 
   382     ||>> lift_thms qtys [] [raw_induct_thm]
   469       |> lift_thms qtys [] raw_size_eqvt
   383     ||>> lift_thms qtys [] raw_exhaust_thms
   470       ||>> lift_thms qtys [] [raw_induct_thm]
   384     ||>> lift_thms qtys [] raw_size_thms
   471       ||>> lift_thms qtys [] raw_exhaust_thms
   385     ||>> lift_thms qtys [] raw_perm_bn_simps
   472       ||>> lift_thms qtys [] raw_size_thms
   386     ||>> lift_thms qtys [] alpha_refl_thms
   473       ||>> lift_thms qtys [] raw_perm_bn_simps
       
   474       ||>> lift_thms qtys [] alpha_refl_thms
       
   475     else raise TEST lthyA
       
   476 
   387 
   477   val qinducts = Project_Rule.projections lthyA qinduct
   388   val qinducts = Project_Rule.projections lthyA qinduct
   478 
   389 
   479   (* supports lemmas *) 
   390   val _ = trace_msg (K "Proving supp lemmas and fs-instances...")
   480   val _ = warning "Proving Supports Lemmas and fs-Instances"
       
   481   val qsupports_thms =
   391   val qsupports_thms =
   482     if get_STEPS lthy > 29
   392     prove_supports lthyB qperm_simps (flat qtrms)
   483     then prove_supports lthyB qperm_simps (flat qtrms)
       
   484     else raise TEST lthyB
       
   485 
   393 
   486   (* finite supp lemmas *)
   394   (* finite supp lemmas *)
   487   val qfsupp_thms =
   395   val qfsupp_thms = prove_fsupp lthyB qtys qinduct qsupports_thms
   488     if get_STEPS lthy > 30
       
   489     then prove_fsupp lthyB qtys qinduct qsupports_thms
       
   490     else raise TEST lthyB
       
   491 
   396 
   492   (* fs instances *)
   397   (* fs instances *)
   493   val lthyC =
   398   val lthyC = fs_instance qtys qty_full_names tvs qfsupp_thms lthyB
   494     if get_STEPS lthy > 31
   399 
   495     then fs_instance qtys qty_full_names tvs qfsupp_thms lthyB
   400   val _ = trace_msg (K "Proving equality between fv and supp...")
   496     else raise TEST lthyB
       
   497 
       
   498   (* fv - supp equality *)
       
   499   val _ = warning "Proving Equality between fv and supp"
       
   500   val qfv_supp_thms = 
   401   val qfv_supp_thms = 
   501     if get_STEPS lthy > 32
   402     prove_fv_supp qtys (flat qtrms) qfvs qfv_bns qalpha_bns qfv_defs qeq_iffs 
   502     then prove_fv_supp qtys (flat qtrms) qfvs qfv_bns qalpha_bns qfv_defs qeq_iffs 
       
   503       qperm_simps qfv_qbn_eqvts qinduct (flat raw_bclauses) lthyC
   403       qperm_simps qfv_qbn_eqvts qinduct (flat raw_bclauses) lthyC
   504     else []
       
   505 
   404 
   506   (* postprocessing of eq and fv theorems *)
   405   (* postprocessing of eq and fv theorems *)
   507 
   406 
   508   val qeq_iffs' = qeq_iffs
   407   val qeq_iffs' = qeq_iffs
   509     |> map (simplify (HOL_basic_ss addsimps qfv_supp_thms))
   408     |> map (simplify (HOL_basic_ss addsimps qfv_supp_thms))
   522   val qfresh_constrs = qsupp_constrs
   421   val qfresh_constrs = qsupp_constrs
   523     |> map (fn thm => thm RS transform_thm) 
   422     |> map (fn thm => thm RS transform_thm) 
   524     |> map (simplify (HOL_basic_ss addsimps transform_thms))
   423     |> map (simplify (HOL_basic_ss addsimps transform_thms))
   525 
   424 
   526   (* proving that the qbn result is finite *)
   425   (* proving that the qbn result is finite *)
   527   val qbn_finite_thms =
   426   val qbn_finite_thms = prove_bns_finite qtys qbns qinduct qbn_defs lthyC
   528     if get_STEPS lthy > 33
       
   529     then prove_bns_finite qtys qbns qinduct qbn_defs lthyC
       
   530     else []
       
   531 
   427 
   532   (* proving that perm_bns preserve alpha *)
   428   (* proving that perm_bns preserve alpha *)
   533   val qperm_bn_alpha_thms = 
   429   val qperm_bn_alpha_thms = 
   534     if get_STEPS lthy > 33
   430     prove_perm_bn_alpha_thms qtys qperm_bns qalpha_bns qinduct qperm_bn_simps qeq_iffs' 
   535     then prove_perm_bn_alpha_thms qtys qperm_bns qalpha_bns qinduct qperm_bn_simps qeq_iffs' 
       
   536       qalpha_refl_thms lthyC
   431       qalpha_refl_thms lthyC
   537     else []
       
   538 
   432 
   539   (* proving the relationship of bn and permute_bn *)
   433   (* proving the relationship of bn and permute_bn *)
   540   val qpermute_bn_thms = 
   434   val qpermute_bn_thms = 
   541     if get_STEPS lthy > 33
   435     prove_permute_bn_thms qtys qbns qperm_bns qinduct qperm_bn_simps qbn_defs qfv_qbn_eqvts lthyC
   542     then prove_permute_bn_thms qtys qbns qperm_bns qinduct qperm_bn_simps qbn_defs qfv_qbn_eqvts lthyC
   436 
   543     else []
   437   val _ = trace_msg (K "Proving strong exhaust lemmas...")
   544 
       
   545   (* proving the strong exhausts and induct lemmas *)
       
   546   val qstrong_exhaust_thms = prove_strong_exhausts lthyC qexhausts bclauses qbn_finite_thms qeq_iffs'
   438   val qstrong_exhaust_thms = prove_strong_exhausts lthyC qexhausts bclauses qbn_finite_thms qeq_iffs'
   547     qfv_qbn_eqvts qpermute_bn_thms qperm_bn_alpha_thms
   439     qfv_qbn_eqvts qpermute_bn_thms qperm_bn_alpha_thms
   548 
   440 
       
   441   val _ = trace_msg (K "Proving strong induct lemmas...")
   549   val qstrong_induct_thms =  prove_strong_induct lthyC qinduct qstrong_exhaust_thms qsize_simps bclauses
   442   val qstrong_induct_thms =  prove_strong_induct lthyC qinduct qstrong_exhaust_thms qsize_simps bclauses
   550 
       
   551 
   443 
   552   (* noting the theorems *)  
   444   (* noting the theorems *)  
   553 
   445 
   554   (* generating the prefix for the theorem names *)
   446   (* generating the prefix for the theorem names *)
   555   val thms_name = 
   447   val thms_name = 
   580      ||>> Local_Theory.note ((thms_suffix "bn_finite", []), qbn_finite_thms)
   472      ||>> Local_Theory.note ((thms_suffix "bn_finite", []), qbn_finite_thms)
   581      ||>> Local_Theory.note ((thms_suffix "perm_bn_alpha", []), qperm_bn_alpha_thms)
   473      ||>> Local_Theory.note ((thms_suffix "perm_bn_alpha", []), qperm_bn_alpha_thms)
   582      ||>> Local_Theory.note ((thms_suffix "permute_bn", []), qpermute_bn_thms)
   474      ||>> Local_Theory.note ((thms_suffix "permute_bn", []), qpermute_bn_thms)
   583 in
   475 in
   584   lthy9'
   476   lthy9'
   585 end handle TEST ctxt => ctxt
   477 end 
   586 *}
   478 *}
   587 
   479 
   588 
   480 
   589 section {* Preparing and parsing of the specification *}
   481 section {* Preparing and parsing of the specification *}
   590 
   482 
   806 (* Command Keyword *)
   698 (* Command Keyword *)
   807 val _ = Outer_Syntax.local_theory "nominal_datatype" "test" Keyword.thy_decl
   699 val _ = Outer_Syntax.local_theory "nominal_datatype" "test" Keyword.thy_decl
   808   (main_parser >> nominal_datatype2_cmd)
   700   (main_parser >> nominal_datatype2_cmd)
   809 *}
   701 *}
   810 
   702 
   811 
   703 ML {*
   812 end
   704 trace := true
   813 
   705 *}
   814 
   706 
   815 
   707 
       
   708 end
       
   709 
       
   710 
       
   711