replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
authorboehmes
Thu, 08 Jun 2017 23:37:01 +0200
changeset 66035 de6cd60b1226
parent 66034 ded1c636aece
child 66045 f8c4442bb3a9
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
src/HOL/Int.thy
src/HOL/Tools/lin_arith.ML
src/Provers/Arith/fast_lin_arith.ML
--- a/src/HOL/Int.thy	Wed Jun 07 23:23:48 2017 +0200
+++ b/src/HOL/Int.thy	Thu Jun 08 23:37:01 2017 +0200
@@ -1713,11 +1713,7 @@
   "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
   "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
   "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
-          apply (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM)
-        apply (simp_all only: algebra_simps minus_diff_eq)
-  apply (simp_all only: add.commute [of _ "- (numeral n + numeral n)"])
-  apply (simp_all only: minus_add add.assoc left_minus)
-  done
+  by (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM)
 
 text \<open>Implementations.\<close>
 
--- a/src/HOL/Tools/lin_arith.ML	Wed Jun 07 23:23:48 2017 +0200
+++ b/src/HOL/Tools/lin_arith.ML	Thu Jun 08 23:37:01 2017 +0200
@@ -243,7 +243,7 @@
 fun allows_lin_arith thy (discrete : string list) (U as Type (D, [])) : bool * bool =
       if of_lin_arith_sort thy U then (true, member (op =) discrete D)
       else if member (op =) discrete D then (true, true) else (false, false)
-  | allows_lin_arith sg discrete U = (of_lin_arith_sort sg U, false);
+  | allows_lin_arith sg _ U = (of_lin_arith_sort sg U, false);
 
 fun decomp_typecheck thy (discrete, inj_consts) (T : typ, xxx) : decomp option =
   case T of
@@ -266,7 +266,7 @@
   | decomp_negation thy data
       ((Const (@{const_name Trueprop}, _)) $ (Const (@{const_name Not}, _) $ (Const (rel, T) $ lhs $ rhs))) =
       negate (decomp_typecheck thy data (T, (rel, lhs, rhs)))
-  | decomp_negation thy data _ =
+  | decomp_negation _ _ _ =
       NONE;
 
 fun decomp ctxt : term -> decomp option =
@@ -280,6 +280,75 @@
   | domain_is_nat _ = false;
 
 
+(* Abstraction of terms *)
+
+(*
+  Abstract terms contain only arithmetic operators and relations.
+
+  When constructing an abstract term for an arbitrary term, non-arithmetic sub-terms
+  are replaced by fresh variables which are declared in the context. Constructing
+  an abstract term from an arbitrary term follows the strategy of decomp.
+*)
+
+fun apply t u = t $ u
+
+fun with2 f c t u cx = f t cx ||>> f u |>> (fn (t, u) => c $ t $ u)
+
+fun abstract_atom (t as Free _) cx = (t, cx)
+  | abstract_atom (t as Const _) cx = (t, cx)
+  | abstract_atom t (cx as (terms, ctxt)) =
+      (case AList.lookup Envir.aeconv terms t of
+        SOME u => (u, cx)
+      | NONE =>
+          let
+            val (n, ctxt') = yield_singleton Variable.variant_fixes "" ctxt
+            val u = Free (n, fastype_of t)
+          in (u, ((t, u) :: terms, ctxt')) end)
+
+fun abstract_num t cx = if can HOLogic.dest_number t then (t, cx) else abstract_atom t cx
+
+fun is_field_sort (_, ctxt) T = of_field_sort (Proof_Context.theory_of ctxt) (domain_type T)
+
+fun is_inj_const (_, ctxt) f =
+  let val {inj_consts, ...} = get_arith_data ctxt
+  in member (op =) inj_consts f end
+
+fun abstract_arith ((c as Const (@{const_name Groups.plus}, _)) $ u1 $ u2) cx =
+      with2 abstract_arith c u1 u2 cx
+  | abstract_arith (t as (c as Const (@{const_name Groups.minus}, T)) $ u1 $ u2) cx =
+      if nT T then abstract_atom t cx else with2 abstract_arith c u1 u2 cx
+  | abstract_arith (t as (c as Const (@{const_name Groups.uminus}, T)) $ u) cx =
+      if nT T then abstract_atom t cx else abstract_arith u cx |>> apply c
+  | abstract_arith ((c as Const (@{const_name Suc}, _)) $ u) cx = abstract_arith u cx |>> apply c
+  | abstract_arith ((c as Const (@{const_name Groups.times}, _)) $ u1 $ u2) cx =
+      with2 abstract_arith c u1 u2 cx
+  | abstract_arith (t as (c as Const (@{const_name Rings.divide}, T)) $ u1 $ u2) cx =
+      if is_field_sort cx T then with2 abstract_arith c u1 u2 cx else abstract_atom t cx
+  | abstract_arith (t as (c as Const f) $ u) cx =
+      if is_inj_const cx f then abstract_arith u cx |>> apply c else abstract_num t cx
+  | abstract_arith t cx = abstract_num t cx
+
+fun is_lin_arith_rel @{const_name Orderings.less} = true
+  | is_lin_arith_rel @{const_name Orderings.less_eq} = true
+  | is_lin_arith_rel @{const_name HOL.eq} = true
+  | is_lin_arith_rel _ = false
+
+fun is_lin_arith_type (_, ctxt) T =
+  let val {discrete, ...} = get_arith_data ctxt
+  in fst (allows_lin_arith (Proof_Context.theory_of ctxt) discrete T) end
+
+fun abstract_rel (t as (r as Const (rel, Type ("fun", [U, _]))) $ lhs $ rhs) cx =
+      if is_lin_arith_rel rel andalso is_lin_arith_type cx U then with2 abstract_arith r lhs rhs cx
+      else abstract_atom t cx
+  | abstract_rel t cx = abstract_atom t cx
+
+fun abstract_neg ((c as Const (@{const_name Not}, _)) $ t) cx = abstract_rel t cx |>> apply c
+  | abstract_neg t cx = abstract_rel t cx
+
+fun abstract ((c as Const (@{const_name Trueprop}, _)) $ t) cx = abstract_neg t cx |>> apply c
+  | abstract t cx = abstract_atom t cx
+
+
 (*---------------------------------------------------------------------------*)
 (* the following code performs splitting of certain constants (e.g., min,    *)
 (* max) in a linear arithmetic problem; similar to what split_tac later does *)
--- a/src/Provers/Arith/fast_lin_arith.ML	Wed Jun 07 23:23:48 2017 +0200
+++ b/src/Provers/Arith/fast_lin_arith.ML	Thu Jun 08 23:37:01 2017 +0200
@@ -49,6 +49,12 @@
   val decomp: Proof.context -> term -> decomp option
   val domain_is_nat: term -> bool
 
+  (*abstraction for proof replay*)
+  val abstract_arith: term -> (term * term) list * Proof.context ->
+    term * ((term * term) list * Proof.context)
+  val abstract: term -> (term * term) list * Proof.context ->
+    term * ((term * term) list * Proof.context)
+
   (*preprocessing, performed on a representation of subgoals as list of premises:*)
   val pre_decomp: Proof.context -> typ list * term list -> (typ list * term list) list
 
@@ -288,7 +294,7 @@
 fun extract_first p =
   let
     fun extract xs (y::ys) = if p y then (y, xs @ ys) else extract (y::xs) ys
-      | extract xs [] = raise List.Empty
+      | extract _ [] = raise List.Empty
   in extract [] end;
 
 fun print_ineqs ctxt ineqs =
@@ -373,7 +379,7 @@
 with 0 <= n.
 *)
 local
-  exception FalseE of thm
+  exception FalseE of thm * (int * cterm) list * Proof.context
 in
 
 fun mkthm ctxt asms (just: injust) =
@@ -439,29 +445,53 @@
                  | THM _ => mult_by_add n thm
           end);
 
-    fun mult_thm (n, thm) =
+    fun mult_thm n thm =
       if n = ~1 then thm RS LA_Logic.sym
       else if n < 0 then mult (~n) thm RS LA_Logic.sym
       else mult n thm;
 
-    fun simp thm =
+    fun simp thm (cx as (_, hyps, ctxt')) =
       let val thm' = trace_thm ctxt ["Simplified:"] (full_simplify simpset_ctxt thm)
-      in if LA_Logic.is_False thm' then raise FalseE thm' else thm' end;
+      in if LA_Logic.is_False thm' then raise FalseE (thm', hyps, ctxt') else (thm', cx) end;
+
+    fun abs_thm i (cx as (terms, hyps, ctxt)) =
+      (case AList.lookup (op =) hyps i of
+        SOME ct => (Thm.assume ct, cx)
+      | NONE =>
+          let
+            val thm = nth asms i
+            val (t, (terms', ctxt')) = LA_Data.abstract (Thm.prop_of thm) (terms, ctxt)
+            val ct = Thm.cterm_of ctxt' t
+          in (Thm.assume ct, (terms', (i, ct) :: hyps, ctxt')) end);
+
+    fun nat_thm t (terms, hyps, ctxt) =
+      let val (t', (terms', ctxt')) = LA_Data.abstract_arith t (terms, ctxt)
+      in (LA_Logic.mk_nat_thm thy t', (terms', hyps, ctxt')) end;
 
-    fun mk (Asm i) = trace_thm ctxt ["Asm " ^ string_of_int i] (nth asms i)
-      | mk (Nat i) = trace_thm ctxt ["Nat " ^ string_of_int i] (LA_Logic.mk_nat_thm thy (nth atoms i))
-      | mk (LessD j) = trace_thm ctxt ["L"] (hd ([mk j] RL lessD))
-      | mk (NotLeD j) = trace_thm ctxt ["NLe"] (mk j RS LA_Logic.not_leD)
-      | mk (NotLeDD j) = trace_thm ctxt ["NLeD"] (hd ([mk j RS LA_Logic.not_leD] RL lessD))
-      | mk (NotLessD j) = trace_thm ctxt ["NL"] (mk j RS LA_Logic.not_lessD)
-      | mk (Added (j1, j2)) = simp (trace_thm ctxt ["+"] (add_thms (mk j1) (mk j2)))
-      | mk (Multiplied (n, j)) =
-          (trace_msg ctxt ("*" ^ string_of_int n); trace_thm ctxt ["*"] (mult_thm (n, mk j)))
+    fun step0 msg (thm, cx) = (trace_thm ctxt [msg] thm, cx)
+    fun step1 msg j f cx = mk j cx |>> f |>> trace_thm ctxt [msg]
+    and step2 msg j1 j2 f cx = mk j1 cx ||>> mk j2 |>> f |>> trace_thm ctxt [msg]
 
+    and mk (Asm i) cx = step0 ("Asm " ^ string_of_int i) (abs_thm i cx)
+      | mk (Nat i) cx = step0 ("Nat " ^ string_of_int i) (nat_thm (nth atoms i) cx)
+      | mk (LessD j) cx = step1 "L" j (fn thm => hd ([thm] RL lessD)) cx
+      | mk (NotLeD j) cx = step1 "NLe" j (fn thm => thm RS LA_Logic.not_leD) cx
+      | mk (NotLeDD j) cx = step1 "NLeD" j (fn thm => hd ([thm RS LA_Logic.not_leD] RL lessD)) cx
+      | mk (NotLessD j) cx = step1 "NL" j (fn thm => thm RS LA_Logic.not_lessD) cx
+      | mk (Added (j1, j2)) cx = step2 "+" j1 j2 (uncurry add_thms) cx |-> simp
+      | mk (Multiplied (n, j)) cx =
+          (trace_msg ctxt ("*" ^ string_of_int n); step1 "*" j (mult_thm n) cx)
+
+    fun finish ctxt' hyps thm =
+      thm
+      |> fold_rev (Thm.implies_intr o snd) hyps
+      |> singleton (Variable.export ctxt' ctxt)
+      |> fold (fn (i, _) => fn thm => nth asms i RS thm) hyps
   in
     let
       val _ = trace_msg ctxt "mkthm";
-      val thm = trace_thm ctxt ["Final thm:"] (mk just);
+      val (thm, (_, hyps, ctxt')) = mk just ([], [], ctxt);
+      val _ = trace_thm ctxt ["Final thm:"] thm;
       val fls = simplify simpset_ctxt thm;
       val _ = trace_thm ctxt ["After simplification:"] fls;
       val _ =
@@ -472,8 +502,9 @@
             ["Proved:", Thm.string_of_thm ctxt fls, ""]));
           warning "Linear arithmetic should have refuted the assumptions.\n\
             \Please inform Tobias Nipkow.")
-    in fls end
-    handle FalseE thm => trace_thm ctxt ["False reached early:"] thm
+    in finish ctxt' hyps fls end
+    handle FalseE (thm, hyps, ctxt') =>
+      trace_thm ctxt ["False reached early:"] (finish ctxt' hyps thm)
   end;
 
 end;
@@ -555,7 +586,7 @@
   fun elim_neq (ineqs : (LA_Data.decomp option * bool) list) :
                (LA_Data.decomp option * bool) list list =
   let
-    fun elim_neq' nat_only ([] : (LA_Data.decomp option * bool) list) :
+    fun elim_neq' _ ([] : (LA_Data.decomp option * bool) list) :
                   (LA_Data.decomp option * bool) list list =
           [[]]
       | elim_neq' nat_only ((NONE, is_nat) :: ineqs) =