667 end |
673 end |
668 and translate_app_default thy algbr funcgr thm (c_ty, ts) = |
674 and translate_app_default thy algbr funcgr thm (c_ty, ts) = |
669 translate_const thy algbr funcgr thm c_ty |
675 translate_const thy algbr funcgr thm c_ty |
670 ##>> fold_map (translate_term thy algbr funcgr thm) ts |
676 ##>> fold_map (translate_term thy algbr funcgr thm) ts |
671 #>> (fn (t, ts) => t `$$ ts) |
677 #>> (fn (t, ts) => t `$$ ts) |
672 and translate_case thy algbr funcgr thm n cases (app as ((c, ty), ts)) = |
678 and translate_case thy algbr funcgr thm (num_args, (t_pos, case_pats)) (c_ty, ts) = |
673 let |
679 let |
674 val (tys, _) = |
680 val (tys, _) = (chop num_args o fst o strip_type o snd) c_ty; |
675 (chop (1 + (if null cases then 1 else length cases)) o fst o strip_type) ty; |
681 val t = nth ts t_pos; |
676 val dt = nth ts n; |
682 val ty = nth tys t_pos; |
677 val dty = nth tys n; |
683 val ts_clause = nth_drop t_pos ts; |
678 fun is_undefined (Const (c, _)) = Code.is_undefined thy c |
684 fun mk_clause (co, num_co_args) t = |
679 | is_undefined _ = false; |
|
680 fun mk_case (co, n) t = |
|
681 let |
685 let |
682 val _ = if (is_some o Code.get_datatype_of_constr thy) co then () |
686 val _ = if (is_some o Code.get_datatype_of_constr thy) co then () |
683 else error ("Non-constructor " ^ quote co |
687 else error ("Non-constructor " ^ quote co |
684 ^ " encountered in case pattern" |
688 ^ " encountered in case pattern" |
685 ^ (case thm of NONE => "" |
689 ^ (case thm of NONE => "" |
686 | SOME thm => ", in equation\n" ^ Display.string_of_thm thm)) |
690 | SOME thm => ", in equation\n" ^ Display.string_of_thm thm)) |
687 val (vs, body) = Term.strip_abs_eta n t; |
691 val (vs, body) = Term.strip_abs_eta num_co_args t; |
688 val selector = list_comb (Const (co, map snd vs ---> dty), map Free vs); |
692 val not_undefined = case body |
689 in if is_undefined body then NONE else SOME (selector, body) end; |
693 of (Const (c, _)) => not (Code.is_undefined thy c) |
690 fun mk_ds [] = |
694 | _ => true; |
|
695 val pat = list_comb (Const (co, map snd vs ---> ty), map Free vs); |
|
696 in (not_undefined, (pat, body)) end; |
|
697 val clauses = if null case_pats then let val ([v_ty], body) = |
|
698 Term.strip_abs_eta 1 (the_single ts_clause) |
|
699 in [(true, (Free v_ty, body))] end |
|
700 else map (uncurry mk_clause) |
|
701 (AList.make (Code_Unit.no_args thy) case_pats ~~ ts_clause); |
|
702 fun retermify ty (_, (IVar x, body)) = |
|
703 (x, ty) `|-> body |
|
704 | retermify _ (_, (pat, body)) = |
691 let |
705 let |
692 val ([v_ty], body) = Term.strip_abs_eta 1 (the_single (nth_drop n ts)) |
706 val (IConst (_, (_, tys)), ts) = unfold_app pat; |
693 in [(Free v_ty, body)] end |
707 val vs = map2 (fn IVar x => fn ty => (x, ty)) ts tys; |
694 | mk_ds cases = map_filter (uncurry mk_case) |
708 in vs `|--> body end; |
695 (AList.make (Code_Unit.no_args thy) cases ~~ nth_drop n ts); |
709 fun mk_icase const t ty clauses = |
|
710 let |
|
711 val (ts1, ts2) = chop t_pos (map (retermify ty) clauses); |
|
712 in |
|
713 ICase (((t, ty), map_filter (fn (b, d) => if b then SOME d else NONE) clauses), |
|
714 const `$$ (ts1 @ t :: ts2)) |
|
715 end; |
696 in |
716 in |
697 translate_term thy algbr funcgr thm dt |
717 translate_const thy algbr funcgr thm c_ty |
698 ##>> translate_typ thy algbr funcgr dty |
718 ##>> translate_term thy algbr funcgr thm t |
699 ##>> fold_map (fn (pat, body) => translate_term thy algbr funcgr thm pat |
719 ##>> translate_typ thy algbr funcgr ty |
700 ##>> translate_term thy algbr funcgr thm body) (mk_ds cases) |
720 ##>> fold_map (fn (b, (pat, body)) => translate_term thy algbr funcgr thm pat |
701 ##>> translate_app_default thy algbr funcgr thm app |
721 ##>> translate_term thy algbr funcgr thm body |
702 #>> (fn (((dt, dty), ds), t0) => ICase (((dt, dty), ds), t0)) |
722 #>> pair b) clauses |
|
723 #>> (fn (((const, t), ty), ds) => mk_icase const t ty ds) |
703 end |
724 end |
704 and translate_app thy algbr funcgr thm ((c, ty), ts) = case Code.get_case_data thy c |
725 and translate_app thy algbr funcgr thm ((c, ty), ts) = case get_case_scheme thy c |
705 of SOME (n, cases) => let val i = 1 + (if null cases then 1 else length cases) in |
726 of SOME (case_scheme as (num_args, _)) => |
706 if length ts < i then |
727 if length ts < num_args then |
707 let |
728 let |
708 val k = length ts; |
729 val k = length ts; |
709 val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty; |
730 val tys = (curry Library.take (num_args - k) o curry Library.drop k o fst o strip_type) ty; |
710 val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context; |
731 val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context; |
711 val vs = Name.names ctxt "a" tys; |
732 val vs = Name.names ctxt "a" tys; |
712 in |
733 in |
713 fold_map (translate_typ thy algbr funcgr) tys |
734 fold_map (translate_typ thy algbr funcgr) tys |
714 ##>> translate_case thy algbr funcgr thm n cases ((c, ty), ts @ map Free vs) |
735 ##>> translate_case thy algbr funcgr thm case_scheme ((c, ty), ts @ map Free vs) |
715 #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t) |
736 #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t) |
716 end |
737 end |
717 else if length ts > i then |
738 else if length ts > num_args then |
718 translate_case thy algbr funcgr thm n cases ((c, ty), Library.take (i, ts)) |
739 translate_case thy algbr funcgr thm case_scheme ((c, ty), Library.take (num_args, ts)) |
719 ##>> fold_map (translate_term thy algbr funcgr thm) (Library.drop (i, ts)) |
740 ##>> fold_map (translate_term thy algbr funcgr thm) (Library.drop (num_args, ts)) |
720 #>> (fn (t, ts) => t `$$ ts) |
741 #>> (fn (t, ts) => t `$$ ts) |
721 else |
742 else |
722 translate_case thy algbr funcgr thm n cases ((c, ty), ts) |
743 translate_case thy algbr funcgr thm case_scheme ((c, ty), ts) |
723 end |
|
724 | NONE => translate_app_default thy algbr funcgr thm ((c, ty), ts); |
744 | NONE => translate_app_default thy algbr funcgr thm ((c, ty), ts); |
725 |
745 |
726 |
746 |
727 (* store *) |
747 (* store *) |
728 |
748 |