src/HOL/Library/code_lazy.ML
changeset 69568 de09a7261120
parent 68549 bbc742358156
child 69593 3dda49e08b9d
equal deleted inserted replaced
69567:6b4c41037649 69568:de09a7261120
   528   end;
   528   end;
   529 
   529 
   530 fun transform_code_eqs _ [] = NONE
   530 fun transform_code_eqs _ [] = NONE
   531   | transform_code_eqs ctxt eqs =
   531   | transform_code_eqs ctxt eqs =
   532     let
   532     let
       
   533       fun replace_ctr ctxt =
       
   534         let 
       
   535           val thy = Proof_Context.theory_of ctxt
       
   536           val table = Laziness_Data.get thy
       
   537         in fn (s1, s2) => case Symtab.lookup table s1 of
       
   538             NONE => false
       
   539           | SOME x => #active x andalso s2 <> (#ctr x |> dest_Const |> fst)
       
   540         end
   533       val thy = Proof_Context.theory_of ctxt
   541       val thy = Proof_Context.theory_of ctxt
   534       val table = Laziness_Data.get thy
   542       val table = Laziness_Data.get thy
   535       fun eliminate (s1, s2) = case Symtab.lookup table s1 of
       
   536           NONE => false
       
   537         | SOME x => #active x andalso s2 <> (#ctr x |> dest_Const |> fst)
       
   538       fun num_consts_fun (_, T) =
   543       fun num_consts_fun (_, T) =
   539         let
   544         let
   540           val s = body_type T |> dest_Type |> fst
   545           val s = body_type T |> dest_Type |> fst
   541         in
   546         in
   542           if Symtab.defined table s
   547           if Symtab.defined table s
   543             then Ctr_Sugar.ctr_sugar_of ctxt s |> the |> #ctrs |> length
   548           then Ctr_Sugar.ctr_sugar_of ctxt s |> the |> #ctrs |> length
   544             else Code.get_type thy s |> fst |> snd |> length
   549           else Code.get_type thy s |> fst |> snd |> length
   545         end
   550         end
   546       val eqs = map (apfst (Thm.transfer thy)) eqs;
   551       val eqs = map (apfst (Thm.transfer thy)) eqs;
   547 
   552 
   548       val ((code_eqs, nbe_eqs), pure) =
   553       val ((code_eqs, nbe_eqs), pure) =
   549         ((case hd eqs |> fst |> Thm.prop_of of
   554         ((case hd eqs |> fst |> Thm.prop_of of
   552          | _ => (eqs, false))
   557          | _ => (eqs, false))
   553         |> apfst (List.partition snd))
   558         |> apfst (List.partition snd))
   554         handle THM _ => (([], eqs), false)
   559         handle THM _ => (([], eqs), false)
   555       val to_original_eq = if pure then map (apfst (fn x => x RS @{thm eq_reflection})) else I
   560       val to_original_eq = if pure then map (apfst (fn x => x RS @{thm eq_reflection})) else I
   556     in
   561     in
   557       case Case_Converter.to_case ctxt eliminate num_consts_fun (map fst code_eqs) of
   562       case Case_Converter.to_case ctxt (Case_Converter.replace_by_type replace_ctr) num_consts_fun (map fst code_eqs) of
   558           NONE => NONE
   563           NONE => NONE
   559         | SOME thms => SOME (nbe_eqs @ map (rpair true) thms |> to_original_eq)
   564         | SOME thms => SOME (nbe_eqs @ map (rpair true) thms |> to_original_eq)
   560     end handle THM ex => (Output.writeln (@{make_string} eqs); raise THM ex);
   565     end
   561 
   566 
   562 val activate_lazy_type = set_active_lazy_type true;
   567 val activate_lazy_type = set_active_lazy_type true;
   563 val deactivate_lazy_type = set_active_lazy_type false;
   568 val deactivate_lazy_type = set_active_lazy_type false;
   564 val activate_lazy_types = set_active_lazy_types true;
   569 val activate_lazy_types = set_active_lazy_types true;
   565 val deactivate_lazy_types = set_active_lazy_types false;
   570 val deactivate_lazy_types = set_active_lazy_types false;