# HG changeset patch # User blanchet # Date 1382340679 -7200 # Node ID 1753c57eb16ce0ae3048ba17fe6cef0203e4e76a # Parent c1daa6e055656c798e8d43df19eac5350f2e38a5 tuning diff -r c1daa6e05565 -r 1753c57eb16c src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML --- a/src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML Mon Oct 21 09:14:05 2013 +0200 +++ b/src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML Mon Oct 21 09:31:19 2013 +0200 @@ -37,8 +37,8 @@ (* TODO: test with sort constraints on As *) (* TODO: use right sorting order for "fp_sort" w.r.t. original BNFs (?) -- treat new variables as deads? *) -fun mutualize_fp_sugars mutualize fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 = - if mutualize orelse has_duplicates (op =) fpTs then +fun mutualize_fp_sugars has_nested fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 = + if has_nested orelse has_duplicates (op =) fpTs then let val thy = Proof_Context.theory_of no_defs_lthy0; @@ -251,13 +251,13 @@ val perm_kks = permute kks; val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts; - val mutualize = exists (fn Type (_, ty_args) => ty_args <> ty_args0) Ts; + val has_nested = exists (fn Type (_, ty_args) => ty_args <> ty_args0) Ts; val perm_callssss = pad_and_indexify_calls perm_fp_sugars0 nn actual_callssss0; val get_perm_indices = map (fn kk => find_index (curry (op =) kk) perm_kks) o get_indices; val ((perm_fp_sugars, fp_sugar_thms), lthy) = - mutualize_fp_sugars mutualize fp perm_bs perm_Ts get_perm_indices perm_callssss + mutualize_fp_sugars has_nested fp perm_bs perm_Ts get_perm_indices perm_callssss perm_fp_sugars0 lthy; val fp_sugars = unpermute perm_fp_sugars; diff -r c1daa6e05565 -r 1753c57eb16c src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML --- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML Mon Oct 21 09:14:05 2013 +0200 +++ b/src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML Mon Oct 21 09:31:19 2013 +0200 @@ -896,7 +896,7 @@ val exclss'' = exclss' |> map (map (fn (idx, t) => (idx, (Option.map (Goal.prove lthy [] [] t #> Thm.close_derivation) (excl_tac idx), t)))); val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss''; - val (obligation_idxss, goalss) = exclss'' + val (goal_idxss, goalss) = exclss'' |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd)) |> split_list o map split_list; @@ -904,7 +904,7 @@ let val def_thms = map (snd o snd) def_thms'; - val exclss' = map (op ~~) (obligation_idxss ~~ thmss'); + val exclss' = map (op ~~) (goal_idxss ~~ thmss'); fun mk_exclsss excls n = (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1)) |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));