1(*  Title:      ZF/arith_data.ML
2    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3
4Arithmetic simplification: cancellation of common terms
5*)
6
7signature ARITH_DATA =
8sig
9  (*the main outcome*)
10  val nat_cancel: simproc list
11  (*tools for use in similar applications*)
12  val gen_trans_tac: Proof.context -> thm -> thm option -> tactic
13  val prove_conv: string -> tactic list -> Proof.context -> thm list -> term * term -> thm option
14  val simplify_meta_eq: thm list -> Proof.context -> thm -> thm
15  (*debugging*)
16  structure EqCancelNumeralsData   : CANCEL_NUMERALS_DATA
17  structure LessCancelNumeralsData : CANCEL_NUMERALS_DATA
18  structure DiffCancelNumeralsData : CANCEL_NUMERALS_DATA
19end;
20
21
22structure ArithData: ARITH_DATA =
23struct
24
25val iT = Ind_Syntax.iT;
26
27val zero = Const(\<^const_name>\<open>zero\<close>, iT);
28val succ = Const(\<^const_name>\<open>succ\<close>, iT --> iT);
29fun mk_succ t = succ $ t;
30val one = mk_succ zero;
31
32val mk_plus = FOLogic.mk_binop \<^const_name>\<open>Arith.add\<close>;
33
34(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
35fun mk_sum []        = zero
36  | mk_sum [t,u]     = mk_plus (t, u)
37  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
38
39(*this version ALWAYS includes a trailing zero*)
40fun long_mk_sum []        = zero
41  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
42
43val dest_plus = FOLogic.dest_bin \<^const_name>\<open>Arith.add\<close> iT;
44
45(* dest_sum *)
46
47fun dest_sum (Const(\<^const_name>\<open>zero\<close>,_)) = []
48  | dest_sum (Const(\<^const_name>\<open>succ\<close>,_) $ t) = one :: dest_sum t
49  | dest_sum (Const(\<^const_name>\<open>Arith.add\<close>,_) $ t $ u) = dest_sum t @ dest_sum u
50  | dest_sum tm = [tm];
51
52(*Apply the given rewrite (if present) just once*)
53fun gen_trans_tac _ _ NONE = all_tac
54  | gen_trans_tac ctxt th2 (SOME th) = ALLGOALS (resolve_tac ctxt [th RS th2]);
55
56(*Use <-> or = depending on the type of t*)
57fun mk_eq_iff(t,u) =
58  if fastype_of t = iT then FOLogic.mk_eq(t,u)
59                       else FOLogic.mk_iff(t,u);
60
61(*We remove equality assumptions because they confuse the simplifier and
62  because only type-checking assumptions are necessary.*)
63fun is_eq_thm th =
64    can FOLogic.dest_eq (FOLogic.dest_Trueprop (Thm.prop_of th));
65
66fun add_chyps chyps ct = Drule.list_implies (map Thm.cprop_of chyps, ct);
67
68fun prove_conv name tacs ctxt prems (t,u) =
69  if t aconv u then NONE
70  else
71  let val prems' = filter_out is_eq_thm prems
72      val goal = Logic.list_implies (map Thm.prop_of prems',
73        FOLogic.mk_Trueprop (mk_eq_iff (t, u)));
74  in SOME (prems' MRS Goal.prove ctxt [] [] goal (K (EVERY tacs)))
75      handle ERROR msg =>
76        (warning (msg ^ "\nCancellation failed: no typing information? (" ^ name ^ ")"); NONE)
77  end;
78
79
80(*** Use CancelNumerals simproc without binary numerals,
81     just for cancellation ***)
82
83val mk_times = FOLogic.mk_binop \<^const_name>\<open>Arith.mult\<close>;
84
85fun mk_prod [] = one
86  | mk_prod [t] = t
87  | mk_prod (t :: ts) = if t = one then mk_prod ts
88                        else mk_times (t, mk_prod ts);
89
90val dest_times = FOLogic.dest_bin \<^const_name>\<open>Arith.mult\<close> iT;
91
92fun dest_prod t =
93      let val (t,u) = dest_times t
94      in  dest_prod t @ dest_prod u  end
95      handle TERM _ => [t];
96
97(*Dummy version: the only arguments are 0 and 1*)
98fun mk_coeff (0, t) = zero
99  | mk_coeff (1, t) = t
100  | mk_coeff _       = raise TERM("mk_coeff", []);
101
102(*Dummy version: the "coefficient" is always 1.
103  In the result, the factors are sorted terms*)
104fun dest_coeff t = (1, mk_prod (sort Term_Ord.term_ord (dest_prod t)));
105
106(*Find first coefficient-term THAT MATCHES u*)
107fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
108  | find_first_coeff past u (t::terms) =
109        let val (n,u') = dest_coeff t
110        in  if u aconv u' then (n, rev past @ terms)
111                          else find_first_coeff (t::past) u terms
112        end
113        handle TERM _ => find_first_coeff (t::past) u terms;
114
115
116(*Simplify #1*n and n*#1 to n*)
117val add_0s = [@{thm add_0_natify}, @{thm add_0_right_natify}];
118val add_succs = [@{thm add_succ}, @{thm add_succ_right}];
119val mult_1s = [@{thm mult_1_natify}, @{thm mult_1_right_natify}];
120val tc_rules = [@{thm natify_in_nat}, @{thm add_type}, @{thm diff_type}, @{thm mult_type}];
121val natifys = [@{thm natify_0}, @{thm natify_ident}, @{thm add_natify1}, @{thm add_natify2},
122               @{thm diff_natify1}, @{thm diff_natify2}];
123
124(*Final simplification: cancel + and **)
125fun simplify_meta_eq rules ctxt =
126  let val ctxt' =
127    put_simpset FOL_ss ctxt
128      delsimps @{thms iff_simps} (*these could erase the whole rule!*)
129      addsimps rules
130      |> fold Simplifier.add_eqcong [@{thm eq_cong2}, @{thm iff_cong2}]
131  in mk_meta_eq o simplify ctxt' end;
132
133val final_rules = add_0s @ mult_1s @ [@{thm mult_0}, @{thm mult_0_right}];
134
135structure CancelNumeralsCommon =
136  struct
137  val mk_sum            = (fn T:typ => mk_sum)
138  val dest_sum          = dest_sum
139  val mk_coeff          = mk_coeff
140  val dest_coeff        = dest_coeff
141  val find_first_coeff  = find_first_coeff []
142
143  val norm_ss1 =
144    simpset_of (put_simpset ZF_ss \<^context> addsimps add_0s @ add_succs @ mult_1s @ @{thms add_ac})
145  val norm_ss2 =
146    simpset_of (put_simpset ZF_ss \<^context> addsimps add_0s @ mult_1s @ @{thms add_ac} @
147      @{thms mult_ac} @ tc_rules @ natifys)
148  fun norm_tac ctxt =
149    ALLGOALS (asm_simp_tac (put_simpset norm_ss1 ctxt))
150    THEN ALLGOALS (asm_simp_tac (put_simpset norm_ss2 ctxt))
151  val numeral_simp_ss =
152    simpset_of (put_simpset ZF_ss \<^context> addsimps add_0s @ tc_rules @ natifys)
153  fun numeral_simp_tac ctxt =
154    ALLGOALS (asm_simp_tac (put_simpset numeral_simp_ss ctxt))
155  val simplify_meta_eq  = simplify_meta_eq final_rules
156  end;
157
158(** The functor argumnets are declared as separate structures
159    so that they can be exported to ease debugging. **)
160
161structure EqCancelNumeralsData =
162  struct
163  open CancelNumeralsCommon
164  val prove_conv = prove_conv "nateq_cancel_numerals"
165  val mk_bal   = FOLogic.mk_eq
166  val dest_bal = FOLogic.dest_eq
167  val bal_add1 = @{thm eq_add_iff [THEN iff_trans]}
168  val bal_add2 = @{thm eq_add_iff [THEN iff_trans]}
169  fun trans_tac ctxt = gen_trans_tac ctxt @{thm iff_trans}
170  end;
171
172structure EqCancelNumerals = CancelNumeralsFun(EqCancelNumeralsData);
173
174structure LessCancelNumeralsData =
175  struct
176  open CancelNumeralsCommon
177  val prove_conv = prove_conv "natless_cancel_numerals"
178  val mk_bal   = FOLogic.mk_binrel \<^const_name>\<open>Ordinal.lt\<close>
179  val dest_bal = FOLogic.dest_bin \<^const_name>\<open>Ordinal.lt\<close> iT
180  val bal_add1 = @{thm less_add_iff [THEN iff_trans]}
181  val bal_add2 = @{thm less_add_iff [THEN iff_trans]}
182  fun trans_tac ctxt = gen_trans_tac ctxt @{thm iff_trans}
183  end;
184
185structure LessCancelNumerals = CancelNumeralsFun(LessCancelNumeralsData);
186
187structure DiffCancelNumeralsData =
188  struct
189  open CancelNumeralsCommon
190  val prove_conv = prove_conv "natdiff_cancel_numerals"
191  val mk_bal   = FOLogic.mk_binop \<^const_name>\<open>Arith.diff\<close>
192  val dest_bal = FOLogic.dest_bin \<^const_name>\<open>Arith.diff\<close> iT
193  val bal_add1 = @{thm diff_add_eq [THEN trans]}
194  val bal_add2 = @{thm diff_add_eq [THEN trans]}
195  fun trans_tac ctxt = gen_trans_tac ctxt @{thm trans}
196  end;
197
198structure DiffCancelNumerals = CancelNumeralsFun(DiffCancelNumeralsData);
199
200
201val nat_cancel =
202 [Simplifier.make_simproc \<^context> "nateq_cancel_numerals"
203   {lhss =
204     [\<^term>\<open>l #+ m = n\<close>, \<^term>\<open>l = m #+ n\<close>,
205      \<^term>\<open>l #* m = n\<close>, \<^term>\<open>l = m #* n\<close>,
206      \<^term>\<open>succ(m) = n\<close>, \<^term>\<open>m = succ(n)\<close>],
207    proc = K EqCancelNumerals.proc},
208  Simplifier.make_simproc \<^context> "natless_cancel_numerals"
209   {lhss =
210     [\<^term>\<open>l #+ m < n\<close>, \<^term>\<open>l < m #+ n\<close>,
211      \<^term>\<open>l #* m < n\<close>, \<^term>\<open>l < m #* n\<close>,
212      \<^term>\<open>succ(m) < n\<close>, \<^term>\<open>m < succ(n)\<close>],
213    proc = K LessCancelNumerals.proc},
214  Simplifier.make_simproc \<^context> "natdiff_cancel_numerals"
215   {lhss =
216     [\<^term>\<open>(l #+ m) #- n\<close>, \<^term>\<open>l #- (m #+ n)\<close>,
217      \<^term>\<open>(l #* m) #- n\<close>, \<^term>\<open>l #- (m #* n)\<close>,
218      \<^term>\<open>succ(m) #- n\<close>, \<^term>\<open>m #- succ(n)\<close>],
219    proc = K DiffCancelNumerals.proc}];
220
221end;
222
223val _ =
224  Theory.setup (Simplifier.map_theory_simpset (fn ctxt =>
225    ctxt addsimprocs ArithData.nat_cancel));
226
227
228(*examples:
229print_depth 22;
230set timing;
231set simp_trace;
232fun test s = (Goal s; by (Asm_simp_tac 1));
233
234test "x #+ y = x #+ z";
235test "y #+ x = x #+ z";
236test "x #+ y #+ z = x #+ z";
237test "y #+ (z #+ x) = z #+ x";
238test "x #+ y #+ z = (z #+ y) #+ (x #+ w)";
239test "x#*y #+ z = (z #+ y) #+ (y#*x #+ w)";
240
241test "x #+ succ(y) = x #+ z";
242test "x #+ succ(y) = succ(z #+ x)";
243test "succ(x) #+ succ(y) #+ z = succ(z #+ y) #+ succ(x #+ w)";
244
245test "(x #+ y) #- (x #+ z) = w";
246test "(y #+ x) #- (x #+ z) = dd";
247test "(x #+ y #+ z) #- (x #+ z) = dd";
248test "(y #+ (z #+ x)) #- (z #+ x) = dd";
249test "(x #+ y #+ z) #- ((z #+ y) #+ (x #+ w)) = dd";
250test "(x#*y #+ z) #- ((z #+ y) #+ (y#*x #+ w)) = dd";
251
252(*BAD occurrence of natify*)
253test "(x #+ succ(y)) #- (x #+ z) = dd";
254
255test "x #* y2 #+ y #* x2 = y #* x2 #+ x #* y2";
256
257test "(x #+ succ(y)) #- (succ(z #+ x)) = dd";
258test "(succ(x) #+ succ(y) #+ z) #- (succ(z #+ y) #+ succ(x #+ w)) = dd";
259
260(*use of typing information*)
261test "x : nat ==> x #+ y = x";
262test "x : nat --> x #+ y = x";
263test "x : nat ==> x #+ y < x";
264test "x : nat ==> x < y#+x";
265test "x : nat ==> x le succ(x)";
266
267(*fails: no typing information isn't visible*)
268test "x #+ y = x";
269
270test "x #+ y < x #+ z";
271test "y #+ x < x #+ z";
272test "x #+ y #+ z < x #+ z";
273test "y #+ z #+ x < x #+ z";
274test "y #+ (z #+ x) < z #+ x";
275test "x #+ y #+ z < (z #+ y) #+ (x #+ w)";
276test "x#*y #+ z < (z #+ y) #+ (y#*x #+ w)";
277
278test "x #+ succ(y) < x #+ z";
279test "x #+ succ(y) < succ(z #+ x)";
280test "succ(x) #+ succ(y) #+ z < succ(z #+ y) #+ succ(x #+ w)";
281
282test "x #+ succ(y) le succ(z #+ x)";
283*)
284