├── .gitattributes
├── .github
└── workflows
│ └── autoinduct.yml
├── .gitignore
├── LICENSE
├── Makefile
├── README.md
├── ac_simplifier
└── elpi
│ ├── Makefile
│ ├── _CoqProject
│ └── theories
│ ├── AC.v
│ ├── Tactic.v
│ └── Tests.v
├── autoinduct
├── Makefile
├── README.md
├── Test.v
├── elpi
│ ├── Makefile
│ ├── _CoqProject
│ └── theories
│ │ ├── Autoinduct.v
│ │ └── Tactic.v
├── ltac1
│ └── Autoinduct.v
├── ltac2
│ └── Autoinduct.v
├── metacoq
│ ├── Makefile
│ ├── _CoqProject
│ └── theories
│ │ ├── Autoinduct.v
│ │ └── Autoinduct_Template.v
├── ocaml
│ ├── .gitignore
│ ├── Makefile
│ ├── _CoqProject
│ ├── src
│ │ ├── autoinduct.ml
│ │ ├── autoinduct.mli
│ │ ├── autoinduct_plugin.mlpack
│ │ └── autoinduction.mlg
│ └── theories
│ │ └── Autoinduct.v
└── opam
├── real_simplifier
├── elpi
│ ├── Makefile
│ ├── Test.v
│ ├── _CoqProject
│ ├── elpi
│ │ ├── reify.elpi
│ │ └── symbols.elpi
│ └── theories
│ │ ├── RealSimplify.v
│ │ ├── Tactic.v
│ │ └── Theory.v
└── ltac2_reflexive
│ └── theories
│ └── RealSimpl.v
└── show
└── elpi
├── Makefile
├── _CoqProject
└── theories
└── Show.v
/.gitattributes:
--------------------------------------------------------------------------------
1 | *.elpi linguist-language=prolog
2 |
--------------------------------------------------------------------------------
/.github/workflows/autoinduct.yml:
--------------------------------------------------------------------------------
1 | name: Test compilation of Autoinduct plugin
2 |
3 | on: [push, pull_request]
4 |
5 | jobs:
6 | build-matrix:
7 | runs-on: ubuntu-latest
8 | strategy:
9 | matrix:
10 | opam_file:
11 | - 'autoinduct/opam'
12 | image:
13 | - 'coqorg/coq:8.17'
14 | fail-fast: false # don't stop jobs if one fails
15 | steps:
16 | - uses: actions/checkout@v2
17 | with:
18 | submodules: true
19 | - uses: coq-community/docker-coq-action@v1
20 | with:
21 | custom_image: ${{ matrix.image }}
22 | opam_file: ${{ matrix.opam_file }}
23 | before_script: |
24 | startGroup "fix permission issues"
25 | sudo chown -R coq:coq .
26 | endGroup
27 | before_install: |
28 | startGroup "Add coq-extra-dev repo"
29 | opam repo add coq-extra-dev https://coq.inria.fr/opam/extra-dev
30 | endGroup
31 | after_script: |
32 | startGroup "Run test"
33 | make test
34 | endGroup
35 |
36 |
--------------------------------------------------------------------------------
/.gitignore:
--------------------------------------------------------------------------------
1 | *.mli~
2 | *.ml~
3 | *.aux
4 | *.glob
5 | *.vo
6 | *.vok
7 | *.vos
8 | *.o
9 | *.cmi
10 | *.cmx
11 | *.cmxs
12 | *.d
13 | *.a
14 | *.cmxa
15 | Makefile.coq
16 | Makefile.coq.conf
17 | **/src/g_tuto*.ml
18 |
--------------------------------------------------------------------------------
/LICENSE:
--------------------------------------------------------------------------------
1 | The MIT License (MIT)
2 |
3 | Permission is hereby granted, free of charge, to any person obtaining a copy
4 | of this software and associated documentation files (the "Software"), to deal
5 | in the Software without restriction, including without limitation the rights
6 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell
7 | copies of the Software, and to permit persons to whom the Software is
8 | furnished to do so, subject to the following conditions:
9 |
10 | The above copyright notice and this permission notice shall be included in
11 | all copies or substantial portions of the Software.
12 |
13 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR
14 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY,
15 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE
16 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER
17 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM,
18 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN
19 | THE SOFTWARE.
20 |
--------------------------------------------------------------------------------
/Makefile:
--------------------------------------------------------------------------------
1 | all: test
2 |
3 | build:
4 | make -C autoinduct build
5 |
6 | test: build
7 | make -C autoinduct test
8 |
9 | install:
10 | echo "This project is not to be installed"
11 | exit 1
12 |
--------------------------------------------------------------------------------
/README.md:
--------------------------------------------------------------------------------
1 | # Rosetta stone of metaprogramming in Coq
2 |
3 | [![Contributing][contributing-shield]][contributing-link]
4 | [![Code of Conduct][conduct-shield]][conduct-link]
5 | [![Zulip][zulip-shield]][zulip-link]
6 |
7 | [contributing-shield]: https://img.shields.io/badge/contributions-welcome-%23f7931e.svg
8 | [contributing-link]: https://github.com/coq-community/manifesto/blob/master/CONTRIBUTING.md
9 |
10 | [conduct-shield]: https://img.shields.io/badge/%E2%9D%A4-code%20of%20conduct-%23f15a24.svg
11 | [conduct-link]: https://github.com/coq-community/manifesto/blob/master/CODE_OF_CONDUCT.md
12 |
13 | [zulip-shield]: https://img.shields.io/badge/chat-on%20zulip-%23c1272d.svg
14 | [zulip-link]: https://coq.zulipchat.com/#narrow/stream/373964-CUDW-2023/topic/Rosetta.20stone.20of.20extension.20languages
15 |
16 | ## Meta
17 |
18 | - Author(s):
19 | - Enzo Crance
20 | - Davide Fissore
21 | - Yannick Forster
22 | - Gaëtan Gilbert
23 | - Talia Ringer
24 | - Michael Soegtrop
25 | - Enrico Tassi
26 | - Tomas Vallejos
27 | - Coq-community maintainer(s):
28 | - Yannick Forster ([**@yforster**](https://github.com/yforster))
29 | - License: [MIT License](LICENSE)
30 | - Compatible Coq versions: 8.17
31 | - Additional dependencies: multiple, see READMEs in directories
32 |
--------------------------------------------------------------------------------
/ac_simplifier/elpi/Makefile:
--------------------------------------------------------------------------------
1 | ifeq "$(COQBIN)" ""
2 | COQBIN=$(dir $(shell which coqtop))/
3 | endif
4 |
5 | %: Makefile.coq
6 |
7 | Makefile.coq: _CoqProject
8 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq
9 |
10 | tests: all
11 | @$(MAKE) -C tests -s clean
12 | @$(MAKE) -C tests -s all
13 |
14 | -include Makefile.coq
15 |
--------------------------------------------------------------------------------
/ac_simplifier/elpi/_CoqProject:
--------------------------------------------------------------------------------
1 | -R theories AC
2 |
3 | theories/AC.v
4 | theories/Tactic.v
5 | theories/Tests.v
6 |
--------------------------------------------------------------------------------
/ac_simplifier/elpi/theories/AC.v:
--------------------------------------------------------------------------------
1 | From Coq Require Import ZArith.
2 |
3 | Fixpoint nth {T : Type} (x0 : T) (s : list T) (n : nat) {struct n} : T :=
4 | match s, n with
5 | | cons x s', S n' => nth x0 s' n'
6 | | cons x _, 0 => x
7 | | _, _ => x0
8 | end.
9 |
10 | Fixpoint ncons {T : Type} (n : nat) (x : T) (s : list T) : list T :=
11 | match n with
12 | | S n' => cons x (ncons n' x s)
13 | | 0 => s
14 | end.
15 |
16 | Fixpoint all {T : Type} (p : T -> bool) (xs : list T) : bool :=
17 | match xs with
18 | | cons x xs' => p x && all p xs'
19 | | _ => true
20 | end.
21 |
22 | (* reified additive group expressions *)
23 | Inductive AGExpr : Set :=
24 | | AGX : nat -> AGExpr (* variable *)
25 | | AGO : AGExpr (* zero *)
26 | | AGOpp : AGExpr -> AGExpr (* opposite *)
27 | | AGAdd : AGExpr -> AGExpr -> AGExpr. (* addition *)
28 |
29 | (* normalizing a reified term to a formal sum, assuming the commutativity of *)
30 | (* addition *)
31 | Fixpoint ZMnorm (e : AGExpr) : list Z :=
32 | match e with
33 | | AGX j => ncons j 0%Z (cons 1%Z nil)
34 | | AGO => nil
35 | | AGOpp e1 => List.map Z.opp (ZMnorm e1)
36 | | AGAdd e1 e2 =>
37 | (fix add_rec (xs ys : list Z) : list Z :=
38 | match xs, ys with
39 | | nil, s | s, nil => s
40 | | cons x xs, cons y ys => cons (Z.add x y) (add_rec xs ys)
41 | end) (ZMnorm e1) (ZMnorm e2)
42 | end.
43 |
44 | Section AGeval.
45 |
46 | Context {G : Type} (zeroG : G) (oppG : G -> G) (addG : G -> G -> G).
47 |
48 | Fixpoint AGeval (vm : list G) (e : AGExpr) : G :=
49 | match e with
50 | | AGX j => nth zeroG vm j
51 | | AGO => zeroG
52 | | AGOpp e1 => oppG (AGeval vm e1)
53 | | AGAdd e1 e2 => addG (AGeval vm e1) (AGeval vm e2)
54 | end.
55 |
56 | Definition mulGz (x : G) (n : Z) : G :=
57 | match n with
58 | | Z0 => zeroG
59 | | Zpos p => Pos.iter (fun y => addG y x) zeroG p
60 | | Zneg p => oppG (Pos.iter (fun y => addG y x) zeroG p)
61 | end.
62 |
63 | Fixpoint ZMsubst (vm : list G) (e : list Z) {struct e} : G :=
64 | match e, vm with
65 | | cons n e', cons x vm' => addG (mulGz x n) (ZMsubst vm' e')
66 | | _, _ => zeroG
67 | end.
68 |
69 | Context (addA : forall x y z : G, addG x (addG y z) = addG (addG x y) z).
70 | Context (addC : forall x y : G, addG x y = addG y x).
71 | Context (add0x : forall x : G, addG zeroG x = x).
72 | Context (addNx : forall x : G, addG (oppG x) x = zeroG).
73 |
74 | Let addx0 x : addG x zeroG = x.
75 | Proof. now rewrite addC, add0x. Qed.
76 |
77 | Let addxN x : addG x (oppG x) = zeroG.
78 | Proof. now rewrite addC, addNx. Qed.
79 |
80 | Let oppK x : oppG (oppG x) = x.
81 | Proof.
82 | now rewrite <- (add0x (oppG (oppG x))), addC, <- (addNx x), addA, addNx, add0x.
83 | Qed.
84 |
85 | Let oppD x y : oppG (addG x y) = addG (oppG x) (oppG y).
86 | Proof.
87 | rewrite <- (add0x (addG (oppG x) _)), <- (addNx (addG y x)), <- addA.
88 | replace (addG (addG y x) (addG (oppG x) (oppG y))) with zeroG.
89 | now rewrite addx0, addC.
90 | now rewrite addA, <- (addA y), addxN, addx0, addxN.
91 | Qed.
92 |
93 | Let oppB x y : oppG (addG x (oppG y)) = addG y (oppG x).
94 | Proof. now rewrite oppD, oppK. Qed.
95 |
96 | Let opp0 : oppG zeroG = zeroG.
97 | Proof. now rewrite <- (add0x (oppG _)), addxN. Qed.
98 |
99 | Let addCA x y z : addG x (addG y z) = addG y (addG x z).
100 | Proof. now rewrite !addA, (addC x). Qed.
101 |
102 | Let addAC x y z : addG (addG x y) z = addG (addG x z) y.
103 | Proof. now rewrite <- !addA, (addC y). Qed.
104 |
105 | Let addACA x y z t : addG (addG x y) (addG z t) = addG (addG x z) (addG y t).
106 | Proof. now rewrite !addA; f_equal; rewrite <- !addA; f_equal; rewrite addC. Qed.
107 |
108 | Let mulzDl x n m : mulGz x (Z.add n m) = addG (mulGz x n) (mulGz x m).
109 | Proof.
110 | assert (Hpos : forall p q,
111 | addG (mulGz x (Z.pos p)) (mulGz x (Z.pos q)) = mulGz x (Z.pos (Pos.add p q))).
112 | intros p q; simpl; rewrite !Pos2Nat.inj_iter, Pos2Nat.inj_add.
113 | induction (Pos.to_nat p) as [|p' IHp']; simpl; trivial.
114 | now rewrite addAC, IHp'.
115 | assert (Hneg : forall p q,
116 | addG (mulGz x (Z.pos p)) (mulGz x (Z.neg q)) = mulGz x (Z.pos p - Z.pos q)).
117 | {
118 | intros p q; change (mulGz x (Z.neg q)) with (oppG (mulGz x (Z.pos q))).
119 | destruct (Pos.compare_spec q p) as [q_eq_p|q_lt_p|p_lt_q].
120 | - now rewrite q_eq_p, addxN, Z.sub_diag.
121 | - rewrite <- (Pos2Z.inj_sub _ _ q_lt_p), <- (addx0 (mulGz x (Z.pos (p - q)))).
122 | rewrite <- (addxN (mulGz x (Z.pos q))), addA; f_equal.
123 | now rewrite Hpos, Pos.sub_add.
124 | - rewrite <- (Z.opp_involutive (_ - _)), Z.opp_sub_distr, Z.add_comm.
125 | rewrite Z.add_opp_r, <- (Pos2Z.inj_sub _ _ p_lt_q).
126 | change (mulGz x (- Z.pos (q - p))) with (oppG (mulGz x (Z.pos (q - p)))).
127 | rewrite <- (add0x (oppG (mulGz x (Z.pos (q - p))))).
128 | rewrite <- (addxN (mulGz x (Z.pos p))), <- addA; f_equal.
129 | now rewrite <- oppD, Hpos, Pplus_minus; [| apply Pos.lt_gt].
130 | }
131 | destruct n as [|n|n], m as [|m|m]; rewrite ?add0x, ?addx0; trivial.
132 | - now rewrite Hpos.
133 | - now rewrite Hneg.
134 | - now rewrite addC, Hneg.
135 | - change (mulGz x (Z.neg n + Z.neg m)) with (oppG (mulGz x (Z.pos (n + m)))).
136 | now rewrite <- Hpos, oppD.
137 | Qed.
138 |
139 | Lemma ZM_norm_subst (vm : list G) (e : AGExpr) :
140 | ZMsubst vm (ZMnorm e) = AGeval vm e.
141 | Proof.
142 | induction e as [j| |e IHe|e1 IHe1 e2 IHe2]; simpl; trivial.
143 | - revert vm; induction j as [|j IHj]; destruct vm as [|v vm]; simpl; trivial.
144 | + now rewrite add0x, addx0.
145 | + now rewrite add0x, IHj.
146 | - rewrite <- IHe; clear IHe; revert vm.
147 | induction (ZMnorm e) as [|x xs IHxs]; destruct vm as [|v vm]; simpl;
148 | try now rewrite opp0.
149 | rewrite IHxs, oppD; f_equal.
150 | now destruct x; simpl; [rewrite opp0| |rewrite oppK].
151 | - rewrite <- IHe1, <- IHe2; clear IHe1 IHe2; revert vm.
152 | generalize (ZMnorm e1) as xs, (ZMnorm e2) as ys; induction xs as [|x xs IHxs];
153 | destruct ys as [|y ys]; destruct vm as [|v vm]; simpl;
154 | try now (rewrite add0x || rewrite addx0).
155 | now rewrite addACA, <-IHxs, mulzDl.
156 | Qed.
157 |
158 | Lemma ZM_correct (vm : list G) (e1 e2 : AGExpr) :
159 | let isZero (n : Z) := match n with Z0 => true | _ => false end in
160 | all isZero (ZMnorm (AGAdd e1 (AGOpp e2))) = true ->
161 | AGeval vm e1 = AGeval vm e2.
162 | Proof.
163 | set (e := AGAdd e1 (AGOpp e2)); intros isZero Hzeros.
164 | rewrite <- (addx0 (AGeval vm e1)), <- (add0x (AGeval vm e2)).
165 | rewrite <- (addNx (AGeval vm e2)), addA at 1; f_equal.
166 | change (addG (AGeval vm e1) (oppG (AGeval vm e2))) with (AGeval vm e).
167 | rewrite <- !ZM_norm_subst; revert vm Hzeros.
168 | induction (ZMnorm e) as [|x xs IHxs]; destruct vm as [|v vm]; simpl; trivial.
169 | now destruct x; try discriminate; rewrite add0x; apply IHxs.
170 | Qed.
171 |
172 | End AGeval.
173 |
174 | Fact ZM_correct_Z (vm : list Z) (e1 e2 : AGExpr) :
175 | all (Z.eqb 0) (ZMnorm (AGAdd e1 (AGOpp e2))) = true ->
176 | AGeval Z0 Z.opp Z.add vm e1 = AGeval Z0 Z.opp Z.add vm e2.
177 | Proof.
178 | now apply (ZM_correct _ _ _ Z.add_assoc Z.add_comm Z.add_0_l Z.add_opp_diag_l).
179 | Qed.
180 |
181 | Strategy expand [AGeval].
182 |
--------------------------------------------------------------------------------
/ac_simplifier/elpi/theories/Tactic.v:
--------------------------------------------------------------------------------
1 | From Coq Require Import ZArith.
2 | From elpi Require Import elpi.
3 | From AC Require Import AC.
4 |
5 | (******************************************************************************)
6 | (* The Z_zmodule tactic *)
7 | (******************************************************************************)
8 |
9 | Ltac Z_zmodule_reflection VM ZE1 ZE2 :=
10 | apply (@ZM_correct_Z VM ZE1 ZE2); [vm_compute; reflexivity].
11 |
12 | Elpi Tactic Z_zmodule.
13 | Elpi Accumulate lp:{{
14 |
15 | pred mem o:list term, o:term, o:term.
16 | mem [X|_] X {{ O }} :- !.
17 | mem [_|XS] X {{ S lp:N }} :- !, mem XS X N.
18 |
19 | pred quote i:term, o:term, o:list term.
20 | quote {{ Z0 }} {{ AGO }} _ :- !.
21 | quote {{ Z.opp lp:In1 }} {{ AGOpp lp:Out1 }} VM :- !, quote In1 Out1 VM.
22 | quote {{ Z.add lp:In1 lp:In2 }} {{ AGAdd lp:Out1 lp:Out2 }} VM :- !,
23 | quote In1 Out1 VM, quote In2 Out2 VM.
24 | quote In {{ AGX lp:N }} VM :- !, mem VM In N.
25 |
26 | pred list-constant o:term, o:list term, o:term.
27 | list-constant T [] {{ @nil lp:T }} :- !.
28 | list-constant T [X|XS] {{ @cons lp:T lp:X lp:XS' }} :- list-constant T XS XS'.
29 |
30 | pred solve i:goal, o:list sealed-goal.
31 | solve (goal _ _ {{ @eq Z lp:T1 lp:T2 }} _ _ as Goal) GS :- !,
32 | quote T1 ZE1 VM, !,
33 | quote T2 ZE2 VM, !,
34 | list-constant {{ Z }} VM VM', !,
35 | zmod-reflection VM' ZE1 ZE2 Goal GS.
36 | solve _ _ :- coq.ltac.fail 0 "The goal is not an equation".
37 |
38 | pred zmod-reflection i:term, i:term, i:term, i:goal, o:list sealed-goal.
39 | zmod-reflection VM ZE1 ZE2 G GS :-
40 | coq.ltac.call "Z_zmodule_reflection" [trm VM, trm ZE1, trm ZE2] G GS.
41 | zmod-reflection _ _ _ _ _ :-
42 | coq.ltac.fail 0 "Not a valid Z-module equation".
43 |
44 | }}.
45 | Elpi Typecheck.
46 |
47 | Tactic Notation "Z_zmodule" := (elpi Z_zmodule).
48 |
--------------------------------------------------------------------------------
/ac_simplifier/elpi/theories/Tests.v:
--------------------------------------------------------------------------------
1 | From Coq Require Import ZArith.
2 | From AC Require Import Tactic.
3 |
4 | Goal forall x y z, (x + y + - z = x + - z + y)%Z.
5 | Proof.
6 | intros x y z.
7 | Z_zmodule.
8 | Qed.
9 |
--------------------------------------------------------------------------------
/autoinduct/Makefile:
--------------------------------------------------------------------------------
1 | build: build-elpi build-ltac2 build-ltac1 build-ocaml build-metacoq
2 |
3 | build-elpi:
4 | # building elpi #####################
5 | make -C elpi
6 |
7 | build-ltac2:
8 | # building ltac2 #####################
9 | coqc -Q ltac2 Autoinduct.Ltac2 ltac2/Autoinduct.v
10 |
11 | build-ltac1:
12 | # building ltac1 #####################
13 | coqc -Q ltac1 Autoinduct.Ltac1 ltac1/Autoinduct.v
14 |
15 | build-ocaml:
16 | # building ocaml #####################
17 | make -C ocaml
18 |
19 | build-metacoq:
20 | # building metacoq #####################
21 | make -C metacoq
22 |
23 | install: install-elpi install-ocaml install-metacoq
24 |
25 | install-elpi:
26 | # installing elpi #####################
27 | make -C elpi install
28 |
29 | # install-ltac2:
30 | # # installing ltac2 #####################
31 | # coqc -Q ltac2 Ltac2.Autoinduct ltac2/Autoinduct.v
32 |
33 | # install-ltac1:
34 | # # installing ltac1 #####################
35 | # coqc -Q ltac1 Ltac1.Autoinduct ltac1/Autoinduct.v
36 |
37 | install-ocaml:
38 | # installing ocaml #####################
39 | make -C ocaml install
40 |
41 | install-metacoq:
42 | # installing metacoq #####################
43 | make -C metacoq install
44 |
45 | test: test-elpi test-ltac2 test-ltac1 test-ocaml test-metacoq
46 |
47 | test-elpi:
48 | # testing elpi #####################
49 | -coqc -Q elpi/theories Autoinduct.Elpi Test.v
50 |
51 | test-ltac2: build-ltac2
52 | # testing ltac2 #####################
53 | coqc -Q ltac2 Autoinduct.Ltac2 Test.v
54 |
55 | test-ltac1: build-ltac1
56 | # testing ltac1 #####################
57 | coqc -Q ltac1 Autoinduct.Ltac1 Test.v
58 |
59 | test-ocaml:
60 | # testing ocaml #####################
61 | OCAMLPATH=$$OCAMLPATH:$$PWD/ocaml/ coqc -I ocaml/src -Q ocaml/theories Autoinduct.OCaml Test.v
62 |
63 | test-metacoq:
64 | # testing metacoq #####################
65 | coqc -Q metacoq/theories Autoinduct.MetaCoq Test.v
66 |
67 | clean:
68 | make -C elpi clean
69 | rm -f ltac2/Autoinduct.vo
70 | rm -f ltac1/Autoinduct.vo
71 |
--------------------------------------------------------------------------------
/autoinduct/README.md:
--------------------------------------------------------------------------------
1 | # The `autoinduct` tactic
2 |
3 | This exercise amounts to implementing a tactic which performs induction on the right argument of a recursive function.
4 | The tactic was originally used in a course by T. Ringer, more information is [here](https://dependenttyp.es/classes/fa2022/artifacts/12-custom.html).
5 |
6 | ## Steps
7 | 1. `autoinduct on (f a b).` should run `induction` on the recursive argument of `f`, in this case either `a` or `b`.
8 | 1. `f` unfolds to a `fix`, like `Nat.plus` does
9 | 1. `f` unfolds to `fun ... => fix`, like `List.map does`
10 | 1. `f` reduces (many steps) to the above
11 | 2. `f` takes any number of arguments, not just two
12 | 1. `autoinduct on f.` as above but the arguments on which `induction` is run are not given. The tactic has to inspect the goal and run `induction` on an actual term.
13 | 2. `autoinduct` as above but finds the first `f` in the goal for which we can run an induction
14 |
15 | The tactic should be callable from the `"Classic"` proof mode, AKA ltac1.
16 |
17 | # Example
18 |
19 | ```coq
20 | Fixpoint add_left (n m : nat) : nat :=
21 | match n with
22 | | O => m
23 | | S p => S (add_left p m)
24 | end.
25 |
26 | Lemma add_left_O : forall (n : nat), add_left n O = n.
27 | Proof.
28 | intros.
29 | autoinduct on (add_left n O).
30 | all: (simpl; congruence).
31 | Qed.
32 | ```
33 |
34 | # Solutions
35 |
36 | ## OCaml
37 |
38 | The code is in [this file](ocaml/src/autoinduct.ml)
39 |
40 |
41 |
42 | expand
43 |
44 | Details specific to the OCaml code.
45 |
46 | The current version:
47 | - does not go under binders, and
48 | - supports all of part 0, part 1, and part 2.
49 |
50 |
51 |
52 | ## LTac1
53 |
54 | The code is in [this file](ltac/Ltac1.v)
55 |
56 | Requires: [StructTact](https://github.com/uwplse/StructTact)
57 |
58 |
59 |
60 |
61 | expand
62 |
63 | Details specific to the Ltac1 code.
64 |
65 | The current version supports only part 1, not parts 0 or part 2.
66 | This is mostly because of time limitations.
67 |
68 | About extracting the recursive argument:
69 | - the match construct lets one access the recursive argument `n` of a fix
70 | as in `fix f _ _ {struct n} := _ end`, but does not support multiple arities.
71 | Hence one needs to provide multiple patterns, eg `fix f _ _ _ {struct n} := _ end`
72 | for ternary functions, and so on.
73 | - this requirement could be loosened for part 0, but not in any easily apparent way for parts 1 or 2.
74 |
75 |
76 |
77 | ## LTac2
78 |
79 | The code is in [this file](ltac/Ltac2.v)
80 |
81 |
82 |
83 | expand
84 |
85 | Some details specific to the Ltac2 code.
86 |
87 | About extracting the recursive argument:
88 | - the code uses APIs in the `Unsafe` namespace to access the raw
89 | syntax of terms. This makes the code work for any arity.
90 |
91 | - Ltac2 `eval red` produces non backtrackable errors when the argument
92 | cannot be reduced (eg opaque constant), so in mode 3 this can cause
93 | the tactic to fail incorrectly.
94 |
95 | The current version:
96 | - does not go under binders, and
97 | - supports all of part 0, part 1, and part 2.
98 |
99 |
100 |
101 |
102 | ## Elpi
103 |
104 | The [autoinduct/elpi/](elpi/) directory contains the code of a typical elpi tactic and the file
105 | [Tactic.v](elpi/theories/Tactic.v) the actual implementation of `autoinduct`.
106 |
107 | Requires: `coq-elpi`
108 |
109 |
110 |
111 | expand
112 |
113 | Some details specific to the Elpi code.
114 |
115 | About extracting the recursive argument:
116 | - whilst elpi supports Coq syntax within quotations,
117 | `{{ fix f _ _ {struct N} := _ end }}` does not let one bind `N`, so the code
118 | uses the raw term ast `fix _ _ N _` to extract the index of the recursive
119 | argument
120 | - since we look at the term ast, the code works for any arity of `f`
121 |
122 |
123 |
124 | ## MetaCoq
125 |
126 | The [code](metacoq/theories/Autoinduct.v)
127 |
128 | Requires: `coq-metacoq-template`
129 |
130 |
131 |
132 | expand
133 |
134 | Some details specific to the MetaCoq code.
135 |
136 |
137 |
--------------------------------------------------------------------------------
/autoinduct/Test.v:
--------------------------------------------------------------------------------
1 | From Autoinduct Require Import Autoinduct.
2 |
3 | Set Default Proof Mode "Classic".
4 |
5 | Inductive nat : Set :=
6 | | O : nat
7 | | S : nat -> nat.
8 |
9 | Fixpoint add_left (n m : nat) : nat :=
10 | match n with
11 | | O => m
12 | | S p => S (add_left p m)
13 | end.
14 |
15 | Fixpoint add_right (n m : nat) : nat :=
16 | match m with
17 | | O => n
18 | | S p => S (add_right n p)
19 | end.
20 |
21 | Inductive list (T : Type) : Type :=
22 | | nil : list T
23 | | cons : T -> list T -> list T.
24 |
25 | Fixpoint map (A B : Type) (f : A -> B) (l : list A) : list B :=
26 | match l with
27 | | @nil _ => nil _
28 | | @cons _ a t => @cons B (f a) (map A B f t)
29 | end.
30 |
31 | Fixpoint app (A : Type) (l1 l2 : list A) : list A :=
32 | match l1 with
33 | | @nil _ => l2
34 | | @cons _ a t1 => @cons _ a (@app A t1 l2)
35 | end.
36 |
37 | Definition add_right_alias := add_right.
38 |
39 | Module Step0.
40 |
41 | Lemma add_left_O : forall (n : nat), add_left n O = n.
42 | Proof.
43 | intros.
44 | autoinduct on (add_left n O) || (idtac "step0 test1 fails"; induction n).
45 | all: (simpl; congruence).
46 | Qed.
47 |
48 | Lemma add_right_O : forall (m : nat), add_right O m = m.
49 | Proof.
50 | intros.
51 | autoinduct on (add_right O m) || (idtac "step0 test2 fails"; induction m).
52 | all: (simpl; congruence).
53 | Qed.
54 |
55 | Lemma add_left_S :
56 | forall (n m : nat),
57 | S (add_left n m) = add_left n (S m).
58 | Proof.
59 | intros.
60 | autoinduct on (add_left n m) || (idtac "step0 test3 fails"; induction n).
61 | all: (simpl; congruence).
62 | Qed.
63 |
64 | Lemma add_right_S :
65 | forall (n m : nat),
66 | S (add_right n m) = add_right (S n) m.
67 | Proof.
68 | intros.
69 | autoinduct on (add_right n m) || (idtac "step0 test4 fails"; induction m).
70 | all: (simpl; congruence).
71 | Qed.
72 |
73 | Lemma add_left_comm:
74 | forall (n m : nat),
75 | add_left n m = add_left m n.
76 | Proof.
77 | intros.
78 | autoinduct on (add_left n m) || (idtac "step0 test5 fails"; induction n).
79 | all: simpl.
80 | - symmetry. apply add_left_O.
81 | - (rewrite IHn || rewrite IHn0). apply add_left_S.
82 | Qed.
83 |
84 | Lemma add_right_comm:
85 | forall (n m : nat),
86 | add_right n m = add_right m n.
87 | Proof.
88 | intros.
89 | autoinduct on (add_right n m) || (idtac "step0 test6 fails"; induction m).
90 | all: simpl.
91 | - symmetry. apply add_right_O.
92 | - (rewrite IHm || rewrite IHn0). apply add_right_S.
93 | Qed.
94 |
95 | Lemma map_last_sym :
96 | forall (A B : Type) (f : A -> B) (a : A) (l : list A),
97 | @app _ (@map A B f l) (@cons _ (f a) (@nil _)) = map _ _ f (@app A l (@cons _ a (@nil _))).
98 | Proof.
99 | intros.
100 | autoinduct on (@map A B f l) || (idtac "step0 test7 fails (many arguments)"; induction l).
101 | all: (simpl; congruence).
102 | Qed.
103 |
104 | Lemma add_right_alias_O : forall (m : nat), add_right_alias O m = m.
105 | Proof.
106 | intros.
107 | autoinduct on (add_right_alias O m) || (idtac "step0 test8 fails (aliased constant)"; induction m).
108 | all: (simpl; congruence).
109 | Qed.
110 |
111 | End Step0.
112 |
113 | Module Step1.
114 |
115 | Lemma add_left_O : forall (n : nat), add_left n O = n.
116 | Proof.
117 | intros.
118 | autoinduct on add_left || (idtac "step1 test1 fails"; induction n).
119 | all: (simpl; congruence).
120 | Qed.
121 |
122 | Lemma add_right_O : forall (m : nat), add_right O m = m.
123 | Proof.
124 | intros.
125 | autoinduct on add_right || (idtac "step1 test2 fails"; induction m).
126 | all: (simpl; congruence).
127 | Qed.
128 |
129 | Lemma add_left_S :
130 | forall (n m : nat),
131 | S (add_left n m) = add_left n (S m).
132 | Proof.
133 | intros.
134 | autoinduct on add_left || (idtac "step1 test3 fails"; induction n).
135 | all: (simpl; congruence).
136 | Qed.
137 |
138 | Lemma add_right_S :
139 | forall (n m : nat),
140 | S (add_right n m) = add_right (S n) m.
141 | Proof.
142 | intros.
143 | autoinduct on add_right || (idtac "step1 test4 fails"; induction m).
144 | all: (simpl; congruence).
145 | Qed.
146 |
147 | Lemma add_left_comm:
148 | forall (n m : nat),
149 | add_left n m = add_left m n.
150 | Proof.
151 | intros.
152 | autoinduct on add_left || (idtac "step1 test5 fails"; induction n).
153 | all: simpl.
154 | - symmetry. apply add_left_O.
155 | - (rewrite IHn || rewrite IHn0). apply add_left_S.
156 | Qed.
157 |
158 | Lemma add_right_comm:
159 | forall (n m : nat),
160 | add_right n m = add_right m n.
161 | Proof.
162 | intros.
163 | autoinduct on add_right || (idtac "step1 test6 fails"; induction m).
164 | all: simpl.
165 | - symmetry. apply add_right_O.
166 | - (rewrite IHm || rewrite IHn0). apply add_right_S.
167 | Qed.
168 |
169 | Lemma map_last_sym :
170 | forall (A B : Type) (f : A -> B) (a : A) (l : list A),
171 | @app _ (@map A B f l) (@cons _ (f a) (@nil _)) = @map _ _ f (@app A l (@cons _ a (@nil _))).
172 | Proof.
173 | intros.
174 | autoinduct on map || (idtac "step1 test7 fails (many arguments)"; induction l).
175 | all: (simpl; congruence).
176 | Qed.
177 |
178 | Lemma add_right_alias_O : forall (m : nat), add_right_alias O m = m.
179 | Proof.
180 | intros.
181 | autoinduct on add_right_alias || (idtac "step1 test8 fails (aliased constant)"; induction m).
182 | all: (simpl; congruence).
183 | Qed.
184 |
185 | End Step1.
186 |
187 | Module Step2.
188 |
189 | Lemma add_left_O : forall (n : nat), add_left n O = n.
190 | Proof.
191 | intros.
192 | autoinduct || (idtac "step2 test1 fails"; induction n).
193 | all: (simpl; congruence).
194 | Qed.
195 |
196 | Lemma add_right_O : forall (m : nat), add_right O m = m.
197 | Proof.
198 | intros.
199 | autoinduct || (idtac "step2 test2 fails"; induction m).
200 | all: (simpl; congruence).
201 | Qed.
202 |
203 | Lemma add_left_S :
204 | forall (n m : nat),
205 | S (add_left n m) = add_left n (S m).
206 | Proof.
207 | intros.
208 | autoinduct || (idtac "step2 test3 fails"; induction n).
209 | all: (simpl; congruence).
210 | Qed.
211 |
212 | Lemma add_right_S :
213 | forall (n m : nat),
214 | S (add_right n m) = add_right (S n) m.
215 | Proof.
216 | intros.
217 | autoinduct || (idtac "step2 test4 fails"; induction m).
218 | all: (simpl; congruence).
219 | Qed.
220 |
221 | Lemma add_left_comm:
222 | forall (n m : nat),
223 | add_left n m = add_left m n.
224 | Proof.
225 | intros.
226 | autoinduct || (idtac "step2 test5 fails"; induction n).
227 | all: simpl.
228 | - symmetry. apply add_left_O.
229 | - (rewrite IHn || rewrite IHn0). apply add_left_S.
230 | Qed.
231 |
232 | Lemma add_right_comm:
233 | forall (n m : nat),
234 | add_right n m = add_right m n.
235 | Proof.
236 | intros.
237 | autoinduct || (idtac "step2 test6 fails"; induction m).
238 | all: simpl.
239 | - symmetry. apply add_right_O.
240 | - (rewrite IHm || rewrite IHn0). apply add_right_S.
241 | Qed.
242 |
243 | Lemma map_last_sym :
244 | forall (A B : Type) (f : A -> B) (a : A) (l : list A),
245 | @app _ (@map A B f l) (@cons _ (f a) (@nil _)) = @map _ _ f (@app A l (@cons _ a (@nil _))).
246 | Proof.
247 | intros.
248 | autoinduct || (idtac "step2 test7 fails (many arguments)"; induction l).
249 | all: (simpl; congruence) || (idtac "step2 test7 (many arguments) autoinduct runs, but chooses argument that doesn't finish this proof"; admit).
250 | Admitted.
251 |
252 | Lemma add_right_alias_O : forall (m : nat), add_right_alias O m = m.
253 | Proof.
254 | intros.
255 | autoinduct || (idtac "step2 test8 fails (aliased constant)"; induction m).
256 | all: (simpl; congruence).
257 | Qed.
258 |
259 | End Step2.
260 |
--------------------------------------------------------------------------------
/autoinduct/elpi/Makefile:
--------------------------------------------------------------------------------
1 | ifeq "$(COQBIN)" ""
2 | COQBIN=$(dir $(shell which coqtop))/
3 | endif
4 |
5 | %: Makefile.coq
6 |
7 | Makefile.coq: _CoqProject
8 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq
9 |
10 | tests: all
11 | @$(MAKE) -C tests -s clean
12 | @$(MAKE) -C tests -s all
13 |
14 | -include Makefile.coq
15 |
--------------------------------------------------------------------------------
/autoinduct/elpi/_CoqProject:
--------------------------------------------------------------------------------
1 | -R theories Autoinduct.Elpi
2 |
3 | theories/Tactic.v
4 | theories/Autoinduct.v
5 |
--------------------------------------------------------------------------------
/autoinduct/elpi/theories/Autoinduct.v:
--------------------------------------------------------------------------------
1 | From Autoinduct.Elpi Require Export Tactic.
2 |
--------------------------------------------------------------------------------
/autoinduct/elpi/theories/Tactic.v:
--------------------------------------------------------------------------------
1 | From elpi Require Import elpi.
2 |
3 | (* For some reasons, Coq's builtin tactics cannot be called by their
4 | "name" directly, unless you alias them as we do in the following line *)
5 | Ltac coq_builtin_induction l := induction l.
6 |
7 |
8 | (******************************************************************)
9 | (* Here we implement "autoinduct on (f x y)" *)
10 |
11 | (* This command begins the definition of a tactic *)
12 | Elpi Tactic autoinduct.
13 |
14 | (* This command puts some elpi code in the "autoinduct" tactic.
15 |
16 | The delimiters "lp:{{" and "}}" are used to inject elpi code in the
17 | middle of a Coq file. You need Visual Studio Code or CoqIDE to
18 | strep trough this file interactively.
19 | *)
20 | Elpi Accumulate lp:{{
21 |
22 | % This function takes in unput a term, the body of a recursive function,
23 | % and outputs an integer standing for the position of the recursive argument
24 | %
25 | % This is the signature of the function where "i:" stands for input,
26 | % and "o:" stands for output.
27 | pred index-of-struct-argument i:term, o:int.
28 | %
29 | % The function is made of two rules, one fires whenever the body is
30 | % an immediate fix. Eg
31 | % Definition plus := fix plus n m {struct n} := ... end.
32 | % The fix term constructor contains the number of the recursive argument
33 | index-of-struct-argument (fix _ N _ _) N.
34 | %
35 | % This other rule is for recursive functions that take extra arguments, eg
36 | % Definition map :=
37 | % fun A B (f : A -> B) => fix map (l : list A) {struct l} := ... end.
38 | % In this case we *cross the binder*, make a recursive call, and
39 | % increment the result by one.
40 | %
41 | % Coq terms are represented in HOAS, hence in "(fun Name Ty Bo)"
42 | % the subterm "Bo" is a function binding the argument. If we pass
43 | % a term "t" to it, we obtain the body where the bound variable is
44 | % replaced by "t".
45 | index-of-struct-argument (fun _ _ F) N1 :-
46 | pi x\ % this crafts a fresh name x,
47 | index-of-struct-argument (F x) N, % get the body and recurse
48 | N1 is N + 1.
49 | % the fix term constructor contains the
50 |
51 | % The entry point of tactics is called "solve", it maps a goal
52 | % to a list of subgoals. A goal is a tuple gathering the proof context
53 | % and the statement to prove, among other things. The last element of
54 | % the tuple is that list of arguments passed to the tactic. Here a
55 | % term consisting in the application of a constant F to a list of
56 | % arguments Xs.
57 | solve (goal _ _ _ _ [trm (app [global (const F)|Xs])] as Goal) NewGoals :- !,
58 | % "std.assert! P Msg" is a idion in elpi: it runs P and if it fails
59 | % it aborts the execution printing Msg.
60 | %
61 | % Here we assert that F has a body Bo (axioms for example don't)
62 | std.assert! (coq.env.const-body F (some Bo)) "function is a variable/axiom",
63 | % Then we compute the index of the recursive argument
64 | std.assert! (index-of-struct-argument Bo N) "function is not a fixpoint",
65 | % we then extract the nth argument from Xs
66 | std.assert! (std.nth N Xs RecArg) "not enough arguments",
67 | % finally we pass the ball to Coq's builtin induction tactic
68 | coq.ltac.call "coq_builtin_induction" [trm RecArg] Goal NewGoals.
69 | }}.
70 |
71 | (* Ask elpi to typcheck the code of the current tactic *)
72 | Elpi Typecheck.
73 |
74 |
75 | (* Here we attach an ltac1 concrete syntax to the elpi code above *)
76 | Tactic Notation "autoinduct" "on" constr(t) :=
77 | elpi autoinduct ltac_term:(t).
78 |
79 | (* We are now ready to try our tactic! *)
80 |
81 | (* Mini test *)
82 | Goal forall n m : nat, n + 0 = n.
83 | intros. now autoinduct on (n + 0).
84 | Qed.
85 |
86 | (******************************************************************)
87 | (* The next step is to support "autoinduct on f" *)
88 |
89 | (* We add more code to the same tactic, in particular a new rule
90 | for solve. In this case the argument is not an application, but
91 | just a constant *)
92 | Elpi Accumulate lp:{{
93 | % autoinduct on f
94 | solve (goal _ _ G _ [trm (global (const C) as F)] as Goal) NewGoals :- !, std.do! [
95 | std.assert! (coq.env.const-body C (some Bo)) "function is a variable",
96 | std.assert! (index-of-struct-argument Bo N) "function is not a fixpoint",
97 |
98 | (pi xs l\ fold-map (app [F|xs]) l _ [xs|l]) => fold-map G [] _ RL,
99 | std.rev {std.map RL (xs\ t\ std.nth N xs t)} L,
100 | std.assert! (not (L = [])) "no occurrence of the function found in goal",
101 | std.map L (t\ tac\ tac = coq.ltac.open (coq.ltac.call "coq_builtin_induction" [trm t])) Tactics,
102 | coq.ltac.or Tactics (seal Goal) NewGoals
103 | ].
104 | }}.
105 |
106 |
107 | (******************************************************************)
108 | (* The last step is to support "autoinduct" *)
109 | Elpi Accumulate lp:{{
110 |
111 | pred find-all-induction-candidates i:term, o:list term.
112 | find-all-induction-candidates G Ts :-
113 | (pi c xs l x\ fold-map (app [global (const c)|xs]) l _ [x|l] :- sigma bo n\
114 | coq.env.const-body c (some bo),
115 | index-of-struct-argument bo n,
116 | std.nth n xs x) => fold-map G [] _ Ts.
117 |
118 | % autoinduct
119 | solve (goal _ _ G _ [] as Goal) NewGoals :-
120 | std.map {std.rev {find-all-induction-candidates G} } (t\ tac\
121 | tac = coq.ltac.open (coq.ltac.call "coq_builtin_induction" [trm t]))
122 | Tactics,
123 | coq.ltac.or Tactics (seal Goal) NewGoals.
124 | }}.
125 |
126 | Elpi Typecheck.
127 |
128 | Tactic Notation "autoinduct" := elpi autoinduct.
129 |
--------------------------------------------------------------------------------
/autoinduct/ltac1/Autoinduct.v:
--------------------------------------------------------------------------------
1 | (*
2 | * This is a modification of a tactic implemented for CS 598 on
3 | * proof automation intially.
4 | *
5 | * The biggest limitation of the Ltac approach is that we must limit
6 | * the number of arguments we can apply it to. We can't easily build
7 | * a version of this tactic that works on arbitrarily many arguments.
8 | *
9 | * We'll build on some tactics from StructTact, and also build a new tactic
10 | * in the style of StructTact. This imports the StructTact library so we can do that:
11 | *)
12 | Require Import StructTact.StructTactics.
13 |
14 | (*
15 | * Proofs about add_left and add_right will follow more easily from induction
16 | * over the first and the second arguments, respectively. We will write a tactic
17 | * that does some of this choosing for us automatically! That way, in the style
18 | * of StructTact, if we change a proof that used add_left to instead use add_right,
19 | * we won't have to change the details of our proofs. In that way, we will make
20 | * our proofs more robust.
21 | *)
22 |
23 | Ltac in_reduced f t :=
24 | let f_red := eval red in f in (* reduce f, but leave goal alone *)
25 | ltac:(t f_red). (* run the next tactic on f_red *)
26 |
27 | (*
28 | * The tactic in_reduced_f matches over a goal, finds a function f applied to arguments
29 | * x and y, and then applies tactical t to x, y, and the reduced f.
30 | *
31 | * Here the limitations of Ltac really shine. I stop at 3 arguments
32 | * because I'm impatient and I'm stuck handling each number of arguments
33 | * separately because of Ltac.
34 | *)
35 | Ltac in_reduced_f f t :=
36 | match goal with
37 | | [ |- context [ f ?x ?y ?z ] ] =>
38 | in_reduced f ltac:(t x y z)
39 | | [ |- context [ f ?x ?y ] ] =>
40 | in_reduced f ltac:(t x y)
41 | | [ |- context [ f ?x ] ] =>
42 | in_reduced f ltac:(t x)
43 | end.
44 |
45 | Module V1.
46 |
47 | (*
48 | * The first version of autoinduct makes some extra assumptions, but doesn't rely on
49 | * StructTact or anything else.
50 | *)
51 | Ltac autoinduct1 f :=
52 | in_reduced_f f ltac:(fun x f =>
53 | lazymatch f with
54 | | (fix f n {struct n} := @?body f n) =>
55 | induction x
56 | end).
57 |
58 | Ltac autoinduct2 f :=
59 | in_reduced_f f ltac:(fun x y f =>
60 | lazymatch f with
61 | | (fix f n m {struct m} := @?body f n m) =>
62 | induction y
63 | | (fix f n m {struct n} := @?body f n m) =>
64 | induction x
65 | end).
66 |
67 | Ltac autoinduct3 f :=
68 | in_reduced_f f ltac:(fun x y z f =>
69 | lazymatch f with
70 | | (fix f n m p {struct p} := @?body f n m p) =>
71 | induction z
72 | | (fix f n m p {struct m} := @?body f n m p) =>
73 | induction y
74 | | (fix f n m p {struct n} := @?body f n m p) =>
75 | induction x
76 | end).
77 |
78 | Ltac autoinduct f := first [autoinduct1 f | autoinduct2 f | autoinduct3 f].
79 |
80 | End V1.
81 |
82 | Module V2.
83 |
84 | (*
85 | * Our tactic makes a whole bunch of assumptions. With StructTact we can get rid of
86 | * some of them.
87 | *)
88 |
89 | Ltac autoinduct1 f :=
90 | in_reduced_f f ltac:(fun x f =>
91 | lazymatch f with
92 | | (fix f n {struct n} := @?body f n) =>
93 | try (rememberNonVars x); generalizeEverythingElse x; induction x
94 | end).
95 |
96 | Ltac autoinduct2 f :=
97 | in_reduced_f f ltac:(fun x y f =>
98 | lazymatch f with
99 | | (fix f n m {struct m} := @?body f n m) =>
100 | try (rememberNonVars y); generalizeEverythingElse y; induction y
101 | | (fix f n m {struct n} := @?body f n m) =>
102 | try (rememberNonVars x); generalizeEverythingElse x; induction x
103 | end).
104 |
105 | Ltac autoinduct3 f :=
106 | in_reduced_f f ltac:(fun x y z f =>
107 | lazymatch f with
108 | | (fix f n m p {struct p} := @?body f n m p) =>
109 | try (rememberNonVars z); generalizeEverythingElse z; induction z
110 | | (fix f n m p {struct m} := @?body f n m p) =>
111 | try (rememberNonVars y); generalizeEverythingElse y; induction y
112 | | (fix f n m p {struct n} := @?body f n m p) =>
113 | try (rememberNonVars x); generalizeEverythingElse x; induction x
114 | end).
115 |
116 | Ltac autoinduct f := intros; first [autoinduct1 f | autoinduct2 f | autoinduct3 f].
117 |
118 | End V2.
119 |
120 | Export V1.
121 | Tactic Notation "autoinduct" "on" constr(f) := autoinduct f.
122 |
--------------------------------------------------------------------------------
/autoinduct/ltac2/Autoinduct.v:
--------------------------------------------------------------------------------
1 | Require Import Ltac2.Ltac2.
2 |
3 | (* stdlib fail does enter so has type unit (indeed it doesn't fail if
4 | there are no focused goals) *)
5 | Ltac2 fail s := Control.backtrack_tactic_failure s.
6 |
7 | (* If you don't want to use Unsafe.kind, but can only handle fixpoints with 2 arguments *)
8 | Ltac2 naive_struct_arg (f : constr) : int :=
9 | lazy_match! f with
10 | (* `fix f n m {struct n} := bla` and `fix f n {struct n} := fun m => bla` are the same thing so this handles both arity 1 and arity 2 when the first argument is the structural argument *)
11 | | (fix f n {struct n} := _) => 0
12 | | (fix f n m {struct m} := _) => 1
13 | end.
14 |
15 | Ltac2 rec struct_arg (f : constr) : int :=
16 | match Constr.Unsafe.kind f with
17 | | Constr.Unsafe.Fix structs which _ _ => Array.get structs which
18 | | Constr.Unsafe.Lambda _ body => Int.add 1 (struct_arg body)
19 | | _ => fail "not a fixpoint"
20 | end.
21 |
22 | Ltac2 find_applied (f : (constr * int) option) : constr :=
23 | match! goal with
24 | | [ |- context [ ?t ] ] =>
25 | match Constr.Unsafe.kind t with
26 | | Constr.Unsafe.App g args =>
27 | let farg :=
28 | match f with
29 | | Some (f,farg) =>
30 | if Constr.equal f g then farg
31 | else fail "applies a different function"
32 | | None =>
33 | (* NB: if we encounter a Qed definition, eval red will
34 | fail without backtracking (API limitation of ltac2).
35 | We could work around this by doing the eval red
36 | through ltac1 which does backtrack. *)
37 | if Constr.is_const g then struct_arg (eval red in $g)
38 | else fail "not a constant"
39 | end
40 | in
41 | if Int.le farg (Array.length args)
42 | then Array.get args farg
43 | else fail "not applied enough"
44 | | _ => fail "not an application"
45 | end
46 | end.
47 |
48 | Ltac2 autoinduct0 (f: constr option) : unit :=
49 | let arg :=
50 | match f with
51 | | Some f =>
52 | match Constr.Unsafe.kind f with
53 | | Constr.Unsafe.App f args =>
54 | (* mode 1: induct on an argument from the given term *)
55 | Array.get args (struct_arg (eval red in $f))
56 | | _ =>
57 | (* mode 2: find f applied in the goal and induct on its argument *)
58 | find_applied (Some (f, struct_arg (eval red in $f)))
59 | end
60 | | None =>
61 | (* mode 3: find the first suitable function and argument in the goal *)
62 | find_applied None
63 | end
64 | in
65 | induction $arg.
66 |
67 | Ltac2 Notation "autoinduct" "on" f(constr) := (autoinduct0 (Some f)).
68 |
69 | Ltac2 Notation "autoinduct" := (autoinduct0 None).
70 |
71 | (* called from Ltac2 (ltac2 doesn't expose congruence so go through ltac1 for it) *)
72 | Goal forall n, n + 0 = n.
73 | intros.
74 | Succeed (autoinduct;simpl;ltac1:(congruence)).
75 | autoinduct on Nat.add;simpl;ltac1:(congruence).
76 | Qed.
77 |
78 | Require Import List.
79 |
80 | Goal forall n, n + 0 = n.
81 | Proof.
82 | intros.
83 | autoinduct on (n + 0);simpl;ltac1:(congruence).
84 | Qed.
85 |
86 | Goal forall l : list nat, l ++ nil = l.
87 | Proof.
88 | intros.
89 | autoinduct on app;simpl;ltac1:(congruence).
90 | Qed.
91 | (* called from ltac1 *)
92 |
93 | Ltac autoinduct :=
94 | ltac2:(f |-
95 | let f :=
96 | Option.get (Ltac1.to_constr f)
97 | in
98 | autoinduct0 (Some f)).
99 |
100 | Tactic Notation "autoinduct" "on" constr(x) := autoinduct x.
101 | Tactic Notation "autoinduct" := ltac2:(autoinduct0 None).
102 |
103 | Set Default Proof Mode "Classic".
104 |
105 | Goal forall n, n + 0 = n.
106 | intros.
107 | autoinduct on Nat.add;simpl;congruence.
108 | Qed.
109 |
--------------------------------------------------------------------------------
/autoinduct/metacoq/Makefile:
--------------------------------------------------------------------------------
1 | ifeq "$(COQBIN)" ""
2 | COQBIN=$(dir $(shell which coqtop))/
3 | endif
4 |
5 | %: Makefile.coq
6 |
7 | Makefile.coq: _CoqProject
8 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq
9 |
10 | tests: all
11 | @$(MAKE) -C tests -s clean
12 | @$(MAKE) -C tests -s all
13 |
14 | -include Makefile.coq
15 |
--------------------------------------------------------------------------------
/autoinduct/metacoq/_CoqProject:
--------------------------------------------------------------------------------
1 | -R theories Autoinduct.MetaCoq
2 |
3 | theories/Autoinduct_Template.v
4 | theories/Autoinduct.v
5 |
--------------------------------------------------------------------------------
/autoinduct/metacoq/theories/Autoinduct.v:
--------------------------------------------------------------------------------
1 | (* General imports to work with TemplateMonad *)
2 | From MetaCoq.Template Require Import All.
3 | From MetaCoq.Utils Require Import bytestring.
4 | Require Import List.
5 | Import MCMonadNotation ListNotations.
6 | Open Scope bs.
7 |
8 | Fixpoint decompose_lam_assum (Γ : context) (t : term) : context * term :=
9 | match t with
10 | | tLambda n A B => decompose_lam_assum (Γ ,, vass n A) B
11 | | tLetIn na b bty b' => decompose_prod_assum (Γ ,, vdef na b bty) b'
12 | | _ => (Γ, t)
13 | end.
14 |
15 | Ltac metacoq_run p f :=
16 | run_template_program (t <- tmQuoteRec f ;;
17 | a <- tmEval lazy (p t) ;;
18 | tmUnquote a)
19 | (fun x => let t := eval unfold my_projT2 in (my_projT2 x) in exact t).
20 |
21 | Definition drop_quantification t :=
22 | snd (decompose_prod_assum [] t).
23 |
24 | Definition autoinduct (p : program) : term :=
25 | let (Σ, t) := p in
26 | (* decompose into head and arguments *)
27 | let (hd, args) := decompose_app t in
28 | (* unfold constant *)
29 | let hd' :=
30 | match hd with
31 | | tConst kn _ =>
32 | match lookup_constant Σ kn with
33 | | Some b => match b.(cst_body) with
34 | | Some b => b
35 | | None => tVar "error: constant has no body"
36 | end
37 | | None => tVar "error: constant not declared"
38 | end
39 | | x => x
40 | end in
41 | let (lambdas, rhd) := decompose_lam_assum [] hd' in
42 | let n := List.length lambdas in
43 | match rhd with
44 | | tFix mfix idx =>
45 | match nth_error mfix idx with
46 | | Some f => match nth_error args (n + f.(rarg)) with
47 | | Some x => x
48 | | None => tVar "not enough arguments for induction"
49 | end
50 | | None => tVar "ill-typed fixpoint"
51 | end
52 | | _ => tVar "passed term does not unfold to a fixpoint"
53 | end.
54 |
55 | (* Tactic for Step 1 *)
56 | Tactic Notation "autoinduct1" "on" constr(f) :=
57 | let x := constr:(ltac:(metacoq_run autoinduct f)) in
58 | induction x.
59 |
60 | Lemma test1 : forall n, n + 0 = n.
61 | Proof.
62 | intros.
63 | autoinduct1 on (plus n 0).
64 | all: cbn; congruence.
65 | Qed.
66 |
67 | (** * Step 2 *)
68 |
69 | Definition eq_const (c1 c2: term) : bool :=
70 | match c1, c2 with
71 | | tConst kn1 _, tConst kn2 _ => kn1 == kn2
72 | | _, _ => false
73 | end.
74 |
75 | (* Checks if the applied term of any tApp in args is f *)
76 | Fixpoint find_cnst_in_args (f : term) (args : list term): term :=
77 | match args with
78 | | nil => tVar "error: passed term does not appear in the goal"
79 | | cons a args => match a with
80 | | tApp hd a_args => if eq_const hd f
81 | then a
82 | else find_cnst_in_args f args
83 | | _ => find_cnst_in_args f args
84 | end
85 | end.
86 |
87 | (* From a list of terms returns all the tApp nodes, including the ones appearing as arguments *)
88 | #[bypass_check(guard)]
89 | Fixpoint split_apps (ts : list term) : list term :=
90 | match ts with
91 | | nil => nil term
92 | | (tApp hd args) :: ts' => (tApp hd args) :: (split_apps args) ++ (split_apps ts')
93 | | _ :: ts' => split_apps ts'
94 | end.
95 |
96 | (* Finds if f appears applied in ctx *)
97 | Definition find_app (ctx f: program) : term :=
98 | let (_, f) := f in
99 | let (Σ, t) := ctx in
100 | let (_, goal) := decompose_prod_assum [] t in
101 | match goal with
102 | | tApp hd args => if eq_const hd f
103 | then goal
104 | else find_cnst_in_args f (split_apps (map (fun t => snd (decompose_prod_assum [] t)) args))
105 | | _ => tVar "error: the goal does not have an application"
106 | end.
107 |
108 | (* Tactic for Step 2 *)
109 | Tactic Notation "autoinduct" "on" constr(f) :=
110 | match goal with
111 | | [ |- ?G ] => run_template_program (goal <- tmQuoteRec G;;
112 | fp <- tmQuoteRec f;;
113 | t <- match fp.2 with
114 | | tApp _ _=> tmReturn fp (* f comes with args *)
115 | | _ => app_f <- tmEval lazy (find_app goal fp);; (* f is just the name of the fixpoint *)
116 | tmReturn (goal.1,app_f)
117 | end;;
118 | a <- tmEval lazy (autoinduct t);;
119 | tmUnquote a
120 | )
121 | (fun x =>
122 | let t := eval unfold my_projT2 in (my_projT2 x) in
123 | induction t
124 | )
125 | end.
126 |
127 | (** * Step 3 *)
128 |
129 | (* Looks in the goal applications of fixpoints *)
130 | Definition find_fixpoints (ctx : program) : list term :=
131 | let (Σ, t) := ctx in
132 | let goal := drop_quantification t in
133 | let fixpoint_candidates := match goal with
134 | | tApp hd args => (goal :: (split_apps (map drop_quantification args)))
135 | | _ => [tVar "error: the goal does not have an application"]
136 | end in
137 |
138 | let fixpoints := filter
139 | (fun t=> match t with
140 | | tApp (tConst kn _) _ => (match lookup_constant Σ kn with
141 | | Some b => true
142 | | _ => false
143 | end)
144 | | _ => false end)
145 | fixpoint_candidates
146 | in
147 | fixpoints.
148 |
149 | (* Step 3 *)
150 | Tactic Notation "autoinduct" :=
151 | match goal with
152 | | [ |- ?G ] => run_template_program (goal <- tmQuoteRec G;;
153 | fixes <- tmEval lazy (find_fixpoints goal);;
154 | pfixes <- tmEval lazy (map (fun t=> (goal.1,t)) fixes);;
155 | autoinducts <- tmEval lazy (map autoinduct pfixes);;
156 | (* filtering candidates which are function calls *)
157 | autoinducts <- tmEval lazy (filter (fun t=> match t with | tApp _ _ => false | _ => true end) autoinducts);;
158 | (* todo filter errors in autoinducts *)
159 | let appf := match head autoinducts with
160 | | None => tVar "error: no candidate fixpoints to make induction"
161 | | Some x => x
162 | end in
163 | a <- tmEval lazy appf;;
164 | tmUnquote a
165 |
166 | )
167 | (fun x =>
168 | let t := eval unfold my_projT2 in (my_projT2 x) in
169 | induction t
170 | )
171 | end.
172 |
173 | Lemma test : forall n, n + 0 = n.
174 | Proof.
175 | intros.
176 | autoinduct on (plus n 0).
177 | all: cbn; congruence.
178 | Qed.
179 |
180 | Lemma map_length : forall [A B : Type] (f : A -> B) (l : list A), #|map f l| = #|l|.
181 | Proof.
182 | intros. autoinduct on (map f l); simpl; auto.
183 | Qed.
184 |
185 | (* works without the args *)
186 | Lemma test2 : forall n, n + 0 = n.
187 | Proof.
188 | intros.
189 | autoinduct on plus.
190 | all: cbn; congruence.
191 | Qed.
192 |
193 | (* still works with (plus n 0) *)
194 | Lemma test2_ : forall n, n + 0 = n.
195 | Proof.
196 | intros.
197 | autoinduct on (plus n 0).
198 | all: cbn; congruence.
199 | Qed.
200 |
201 | (* works without the args *)
202 | Lemma map_length2 : forall [A B : Type] (f : A -> B) (l : list A), #|map f l| = #|l|.
203 | Proof.
204 | intros. autoinduct on map; simpl; auto.
205 | Qed.
206 |
207 | (* still works with (map f l) *)
208 | Lemma map_length2_ : forall [A B : Type] (f : A -> B) (l : list A), #|map f l| = #|l|.
209 | Proof.
210 | intros. autoinduct on (map f l) ; simpl; auto.
211 | Qed.
212 |
213 | Lemma test3 : forall n, n + 0 = n.
214 | Proof.
215 | intros.
216 | autoinduct.
217 | all: cbn; congruence.
218 | Qed.
219 |
220 | Lemma map_length3 : forall [A B : Type] (f : A -> B) (l : list A), #|map f l| = #|l|.
221 | Proof.
222 | intros. autoinduct; simpl; auto.
223 | Qed.
224 |
--------------------------------------------------------------------------------
/autoinduct/metacoq/theories/Autoinduct_Template.v:
--------------------------------------------------------------------------------
1 | (* General imports to work with TemplateMonad *)
2 | From MetaCoq.Template Require Import All.
3 | From MetaCoq.Utils Require Import bytestring.
4 | Require Import List.
5 | Import MCMonadNotation ListNotations.
6 | Open Scope bs.
7 |
8 | Definition autoinduct {A} (a : A) : TemplateMonad typed_term :=
9 | a' <- tmEval cbv a ;;
10 | (* get the inductive representation of a *)
11 | t <- tmQuote a' ;;
12 | (* decompose into head and arguments *)
13 | let (hd, args) := decompose_app t in
14 | match hd with
15 | | tFix mfix idx =>
16 | match nth_error mfix idx with
17 | | Some f => match nth_error args (f.(rarg)) with
18 | | Some x => tmUnquote x
19 | | None => tmFail "not enough arguments for induction"
20 | end
21 | | None => tmFail "ill-typed fixpoint"
22 | end
23 | | _ => tmFail "passed term does not unfold to a fixpoint"
24 | end.
25 |
26 | Tactic Notation "autoinduct" constr(f) :=
27 | run_template_program (autoinduct f)
28 | (* get out the term from the pair of type and term, do induction *)
29 | (fun x => let t := eval unfold my_projT2 in (my_projT2 x) in
30 | induction t).
31 |
32 | Lemma test : forall n, n + 0 = n.
33 | Proof.
34 | intros.
35 | autoinduct (plus n 0).
36 | all: cbn; congruence.
37 | Qed.
38 |
39 | Lemma map_length : forall [A B : Type] (f : A -> B) (l : list A), #|map f l| = #|l|.
40 | Proof.
41 | intros. autoinduct (map f l); simpl; auto.
42 | Qed.
43 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/.gitignore:
--------------------------------------------------------------------------------
1 | *.conf
2 | *.d
3 | *.merlin
4 | src/META.autoinduct-plugin
5 | src/*.cmi
6 | src/*.cml
7 | src/*.d
8 | src/*.cmx
9 | src/*.cmxa
10 | src/*.cmt
11 | src/*.cmti
12 | src/*.o
13 | src/*~
14 | src/*.a
15 | src/*.cmxs
16 | src/autoinduction.ml
17 | theories/*.aux
18 | theories/*.vok
19 | theories/*.vos
20 | theories/*.glob
21 | theories/*.vo
22 | Makefile.coq
23 | theories/LocalTests.v
24 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/Makefile:
--------------------------------------------------------------------------------
1 | ifeq "$(COQBIN)" ""
2 | COQBIN=$(dir $(shell which coqtop))/
3 | endif
4 |
5 | %: Makefile.coq
6 |
7 | Makefile.coq: _CoqProject
8 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq
9 |
10 | tests: all
11 | @$(MAKE) -C tests -s clean
12 | @$(MAKE) -C tests -s all
13 |
14 | -include Makefile.coq
15 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/_CoqProject:
--------------------------------------------------------------------------------
1 | -generate-meta-for-package autoinduct-plugin
2 | -R theories Autoinduct.OCaml
3 | -I src
4 |
5 | theories/Autoinduct.v
6 |
7 | src/autoinduct.ml
8 | src/autoinduct.mli
9 | src/autoinduction.mlg
10 | src/autoinduct_plugin.mlpack
11 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/src/autoinduct.ml:
--------------------------------------------------------------------------------
1 | open Proofview
2 | open EConstr
3 | open Environ
4 | open Evd
5 |
6 | (* --- Useful utilities, adapted from coq-plugin-lib --- *)
7 |
8 | (*
9 | * Return a1 if a1_opt is Some a1; otherwise return a2.
10 | * This could be an Option.fold, but that would add OCaml version dependencies.
11 | *)
12 | let get_with_default (a1_opt : 'a option) (a2 : 'a) : 'a =
13 | match a1_opt with
14 | | Some a1 -> a1
15 | | None -> a2
16 |
17 | (* Look up a definition from an environment *)
18 | let lookup_definition (env : Environ.env) (def : constr) (sigma : evar_map) : constr =
19 | match kind sigma def with
20 | | Constr.Const (c, u) ->
21 | begin match constant_value_in env (c, EConstr.Unsafe.to_instance u) with
22 | | v -> EConstr.of_constr v
23 | | exception NotEvaluableConst _ ->
24 | CErrors.user_err (Pp.str "The supplied term is not a transparent constant")
25 | end
26 | | _ -> CErrors.user_err (Pp.str "The supplied term is not a constant")
27 |
28 | (* Equal, but with evar_map tracking for defense against later changes *)
29 | let eequal (trm1 : constr) (trm2 : constr) (sigma : evar_map) : evar_map * bool =
30 | sigma, EConstr.eq_constr sigma trm1 trm2
31 |
32 | (* Pairwise equality between all elements of two arrays *)
33 | let all_eequal (trms1 : constr array) (trms2 : constr array) (sigma : evar_map) : evar_map * bool =
34 | if Array.length trms1 = Array.length trms2 then
35 | List.fold_left2
36 | (fun (sigma, b) trm1 trm2 ->
37 | let sigma, trms_eq = eequal trm1 trm2 sigma in
38 | sigma, b && trms_eq)
39 | (sigma, true)
40 | (Array.to_list trms1)
41 | (Array.to_list trms2)
42 | else
43 | sigma, false
44 |
45 | (* Push a local binding to an environment *)
46 | let push_local ((n, t) : Names.Name.t Context.binder_annot * constr) (env : Environ.env) : Environ.env =
47 | EConstr.push_rel Context.Rel.Declaration.(LocalAssum (n, t)) env
48 |
49 | (* --- Implementation --- *)
50 |
51 | (*
52 | * Get the recursive argument index of the fixpoint defined by f.
53 | * Return None if f is not a fixpoint.
54 | *)
55 | let rec recursive_index (env : Environ.env) (f : constr) (sigma : evar_map) : int option =
56 | match kind sigma f with
57 | | Constr.Fix ((rec_indexes, i), _) ->
58 | (* get the recursive argument *)
59 | Some (Array.get rec_indexes i)
60 | | Constr.Lambda (n, t, b) ->
61 | (* add parameters to the environment and recurse on the body *)
62 | let env_b = push_local (n, t) env in
63 | Option.map (fun n -> 1 + n) (recursive_index env_b b sigma)
64 | | Constr.Const (c, u) ->
65 | (* look up the constant's body and recurse *)
66 | recursive_index env (lookup_definition env f sigma) sigma
67 | | _ ->
68 | (* give up *)
69 | None
70 |
71 | (*
72 | * Get a recursive argument to the appropriate fixpoint to supply to autoinduct.
73 | * Return None if not found.
74 | *)
75 | let recursive_argument (env : Environ.env) (concl : constr) (f_opt : constr option) (args_opt : constr array option) (sigma : evar_map) : evar_map * constr option =
76 | match kind sigma concl with
77 | | Constr.App (g, g_args) ->
78 | let f = get_with_default f_opt g in
79 | let args = get_with_default args_opt g_args in
80 | let sigma, f_eq = eequal f g sigma in
81 | let sigma, args_eq = all_eequal args g_args sigma in
82 | if f_eq && args_eq then
83 | (* the function is the correct applied fixpoint; find the argument *)
84 | let arg_opt =
85 | Option.bind
86 | (recursive_index env f sigma)
87 | (fun arg_no ->
88 | if arg_no < Array.length g_args then
89 | Some (Array.get g_args arg_no)
90 | else
91 | None)
92 | in sigma, arg_opt
93 | else
94 | (* the function is not the correct applied fixpoint; return None *)
95 | sigma, None
96 | | _ ->
97 | (* the goal term does not apply a function at all *)
98 | sigma, None
99 |
100 | (*
101 | * Inner implementation of autoinduct tactic.
102 | *
103 | * The current version:
104 | * 1. does not go under binders,
105 | * 2. supports all parts at once,
106 | * 3. checks exact equality (rather than convertibility) for earlier parts, and
107 | * 4. does not have the best error handling yet.
108 | *
109 | * The options here deal with the flexibility in input format, since we
110 | * handle all three parts at once.
111 | *
112 | * The call to fold_with_binders is found in EConstr in the Coq OCaml API.
113 | * This is a fold over subterms so that we don't need to manually define
114 | * the trivial behavior of other kinds of subterms. The current implementation
115 | * just stops immmediately if under binders (inside a Lambda or Prod) because
116 | * this is easier to start with. If we want to support fancier goals
117 | * with applications under product types, we can change our call to
118 | * fold_with_binders, but notably we will need to shift the found arguments
119 | * since everything here is de Bruijn indexed.
120 | *)
121 | let find_autoinduct (env : Environ.env) (concl : constr) (f_opt : constr option) (args_opt : constr array option) (sigma : evar_map) : evar_map * constr list =
122 | let rec aux (bound : bool) ((sigma, found) : evar_map * constr list) (concl : constr) : evar_map * constr list =
123 | if bound then
124 | (* do not handle binders yet *)
125 | sigma, found
126 | else
127 | (* find any applicable arguments to induct over *)
128 | let sigma, arg_opt = recursive_argument env concl f_opt args_opt sigma in
129 | let found =
130 | match arg_opt with
131 | | Some arg ->
132 | arg :: found
133 | | None ->
134 | found
135 | in fold_with_binders sigma (fun _ -> true) aux bound (sigma, found) concl
136 | in aux false (sigma, []) concl
137 |
138 | (*
139 | * Given the argument to induct over, invoke the induction tactic.
140 | *)
141 | let induct_on (arg : constr) : unit Proofview.tactic =
142 | let dest_arg = Some true, Tactics.ElimOnConstr (fun _ sigma -> sigma, (arg, Tactypes.NoBindings)) in
143 | Tactics.induction_destruct true false ([dest_arg, (None, None), None], None)
144 |
145 | (*
146 | * Find possible arguments to suggest inducting over.
147 | * Try just the first suggested one (which is last in the list of arguments).
148 | * If there are not any good candidates, fail with an error message.
149 | *)
150 | let do_autoinduct (env : Environ.env) (concl : constr) (f_opt : constr option) (args_opt : constr array option) (sigma : evar_map) : unit tactic =
151 | let sigma, induct_args = find_autoinduct env concl f_opt args_opt sigma in
152 | if CList.is_empty induct_args then
153 | Tacticals.tclFAIL (Pp.str "Could not find anything to induct over")
154 | else
155 | tclBIND
156 | (Proofview.Unsafe.tclEVARS sigma)
157 | (fun () -> Tacticals.tclFIRST (List.map induct_on (List.rev induct_args)))
158 |
159 | (*
160 | * Implementation of autoinduct tactic, top level.
161 | * This enters the tactic monad, gets the environment and state, and calls
162 | * the core autoinduct logic, which returns the appropriate tactic.
163 | *)
164 | let autoinduct (f_opt : constr option) (args_opt : constr array option) : unit tactic =
165 | Goal.enter begin fun gl ->
166 | let env = Goal.env gl in
167 | let sigma = Goal.sigma gl in
168 | let concl = Goal.concl gl in
169 | do_autoinduct env concl f_opt args_opt sigma
170 | end
171 |
172 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/src/autoinduct.mli:
--------------------------------------------------------------------------------
1 | (*
2 | * Interface for implementation of autoinduct tactic, which optionally takes
3 | * a fixpoint and/or arguments to look for, and performs induction over an
4 | * argument in the goal corresponding to the recursive argument of that
5 | * fixpoint (or the first applicable fixpoint it finds, if none is supplied).
6 | *)
7 | val autoinduct :
8 | EConstr.t option -> EConstr.t array option -> unit Proofview.tactic
9 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/src/autoinduct_plugin.mlpack:
--------------------------------------------------------------------------------
1 | Autoinduct
2 | Autoinduction
3 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/src/autoinduction.mlg:
--------------------------------------------------------------------------------
1 | DECLARE PLUGIN "autoinduct-plugin.plugin"
2 |
3 | {
4 |
5 | open Ltac_plugin
6 | open Stdarg
7 | open Autoinduct
8 |
9 | }
10 |
11 | (*
12 | * This is where the autoinduct tactic is implemented.
13 | *
14 | * The OCaml version of this tactic is very flexible.
15 | * You can supply the function and all arguments, just the function,
16 | * or nothing at all. The less you supply, the more guessing the tactic
17 | * has to do when it decides where to induct.
18 | *)
19 | TACTIC EXTEND autoinduct_tactic
20 | | [ "autoinduct" "on" "(" constr(f) constr_list(args) ")" ] ->
21 | { autoinduct (Some f) (Some (Array.of_list args)) }
22 | | [ "autoinduct" "on" constr(f) ] ->
23 | { autoinduct (Some f) None }
24 | | [ "autoinduct" ] ->
25 | { autoinduct None None }
26 | END
27 |
--------------------------------------------------------------------------------
/autoinduct/ocaml/theories/Autoinduct.v:
--------------------------------------------------------------------------------
1 | Declare ML Module "autoinduct-plugin.plugin".
2 |
--------------------------------------------------------------------------------
/autoinduct/opam:
--------------------------------------------------------------------------------
1 | opam-version: "2.0"
2 | version: "dev"
3 | maintainer: "yannick.forster@inria.fr"
4 | homepage: "https://github.com/coq-community/metaprogramming-rosetta-stone"
5 | dev-repo: "git+https://github.com/coq-community/metaprogramming-rosetta-stone"
6 | bug-reports: "https://github.com/uds-psl/coq-library-undecidability/issues"
7 | authors: ["Enzo Crance"
8 | "Davide Fissore"
9 | "Yannick Forster"
10 | "Gaëtan Gilbert"
11 | "Talia Ringer"
12 | "Michael Soegtrop"
13 | "Enrico Tassi"
14 | "Tomas Vallejos"]
15 |
16 | license: "MIT"
17 | build: [
18 | [make "-j%{jobs}%"]
19 | ]
20 | install: [
21 | [make "install"]
22 | ]
23 | depends: [
24 | "coq" {>= "8.16"}
25 | "ocaml"
26 | "coq-metacoq-template"
27 | "coq-elpi"
28 | "coq-struct-tact"
29 | ]
30 |
--------------------------------------------------------------------------------
/real_simplifier/elpi/Makefile:
--------------------------------------------------------------------------------
1 | ifeq "$(COQBIN)" ""
2 | COQBIN=$(dir $(shell which coqtop))/
3 | endif
4 |
5 | %: Makefile.coq
6 |
7 | Makefile.coq: _CoqProject
8 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq
9 |
10 | tests: all
11 | @$(MAKE) -C tests -s clean
12 | @$(MAKE) -C tests -s all
13 |
14 | -include Makefile.coq
15 |
--------------------------------------------------------------------------------
/real_simplifier/elpi/Test.v:
--------------------------------------------------------------------------------
1 | From elpi Require Import elpi.
2 |
3 | Require Import Reals.
4 | Require Import ZArith.
5 | Require Import QArith.
6 |
7 | From RealSimplify Require Import RealSimplify Theory.
8 |
9 | Section Test.
10 | Variable x : R.
11 | Open Scope R.
12 |
13 | Goal (x+(Rmax ((3*4+5^2-6)/7) (1/2)) = x + 31/7).
14 | Proof.
15 | real_simplify.
16 | reflexivity.
17 | Qed.
18 |
19 | End Test.
20 |
--------------------------------------------------------------------------------
/real_simplifier/elpi/_CoqProject:
--------------------------------------------------------------------------------
1 | -R theories RealSimplify
2 | -R elpi RealSimplify.Elpi
3 |
4 | theories/Theory.v
5 | theories/Tactic.v
6 | theories/RealSimplify.v
--------------------------------------------------------------------------------
/real_simplifier/elpi/elpi/reify.elpi:
--------------------------------------------------------------------------------
1 | pred reify i:term, o:term.
2 | reify ({{ IZR lp:Z }} as E) R :-
3 | if (zliteral Z) (R = {{ ER_Z lp:Z }}) (R = {{ ER_R lp:E }}).
4 | reify ({{ Q2R lp:Q }} as E) R :-
5 | if (qliteral Q) (R = {{ ER_Q lp:Q }}) (R = {{ ER_R lp:E }}).
6 | reify {{ (- lp:A)%R }} {{ ER_Unary EU_Opp lp:RA }} :- reify A RA.
7 | reify {{ (/ lp:A)%R }} {{ ER_Unary EU_Inv lp:RA }} :- reify A RA.
8 | reify {{ (lp:A + lp:B)%R }} {{ ER_Binary EB_Add lp:RA lp:RB }} :- reify A RA, reify B RB.
9 | reify {{ (lp:A - lp:B)%R }} {{ ER_Binary EB_Sub lp:RA lp:RB }} :- reify A RA, reify B RB.
10 | reify {{ (lp:A * lp:B)%R }} {{ ER_Binary EB_Mul lp:RA lp:RB }} :- reify A RA, reify B RB.
11 | reify {{ (lp:A / lp:B)%R }} {{ ER_Binary EB_Div lp:RA lp:RB }} :- reify A RA, reify B RB.
12 | reify {{ Rmax lp:A lp:B }} {{ ER_Binary EB_Max lp:RA lp:RB }} :- reify A RA, reify B RB.
13 | reify {{ Rmin lp:A lp:B }} {{ ER_Binary EB_Min lp:RA lp:RB }} :- reify A RA, reify B RB.
14 | reify {{ (lp:A ^ lp:B)%R }} {{ ER_Pow lp:RA lp:RB }} :- reify A RA, reify-nat B RB.
15 | reify A {{ ER_R lp:A }}.
16 |
17 | pred reify-nat i:term, o:term.
18 | reify-nat ({{ Z.to_nat lp:Z }} as E) N :-
19 | if (zliteral Z) (N = {{ EN_lit lp:E }}) (N = {{ EN_gen lp:E }}).
20 | reify-nat E N :-
21 | if (computable-nat E) (N = {{ EN_lit lp:E }}) (N = {{ EN_gen lp:E }}).
22 |
23 | pred computable-nat i:term.
24 | computable-nat {{ (lp:A + lp:B)%nat }} :- computable-nat A, computable-nat B.
25 | computable-nat {{ (lp:A - lp:B)%nat }} :- computable-nat A, computable-nat B.
26 | computable-nat {{ (lp:A * lp:B)%nat }} :- computable-nat A, computable-nat B.
27 | computable-nat {{ S lp:A }} :- computable-nat A.
28 | computable-nat {{ O }}.
29 |
30 | pred zliteral i:term.
31 | zliteral {{ Z0 }}.
32 | zliteral {{ Zpos lp:P }} :- pliteral P.
33 | zliteral {{ Zneg lp:P }} :- pliteral P.
34 |
35 | pred pliteral i:term.
36 | pliteral {{ xI lp:P }} :- pliteral P.
37 | pliteral {{ xO lp:P }} :- pliteral P.
38 | pliteral {{ xH }}.
39 |
40 | pred qliteral i:term.
41 | qliteral {{ (lp:N # lp:D)%Q }} :- zliteral N, pliteral D.
42 |
--------------------------------------------------------------------------------
/real_simplifier/elpi/elpi/symbols.elpi:
--------------------------------------------------------------------------------
1 | pred symbols o:list gref.
2 | symbols [
3 | {{:gref interpret_R }},
4 | {{:gref simplify }},
5 | {{:gref cleanup }},
6 | {{:gref interpret_N }},
7 | {{:gref binary_fun }},
8 | {{:gref unary_fun }},
9 | {{:gref binary_fun_q }},
10 | {{:gref unary_fun_q }},
11 | {{:gref binary_check_args }},
12 | {{:gref unary_check_args }},
13 | {{:gref Qnum }},
14 | {{:gref Qden }},
15 | {{:gref Qopp }},
16 | {{:gref Qinv }},
17 | {{:gref Qplus }},
18 | {{:gref Qminus }},
19 | {{:gref Qmult }},
20 | {{:gref Qdiv }},
21 | {{:gref Qpower }},
22 | {{:gref Qcompare }},
23 | {{:gref CompOpp }},
24 | {{:gref Qminmax.Qmax }},
25 | {{:gref Qminmax.Qmin }},
26 | {{:gref GenericMinMax.gmax }},
27 | {{:gref GenericMinMax.gmin }},
28 | {{:gref Qpower_positive }},
29 | {{:gref pow_pos }},
30 | {{:gref Z.opp }},
31 | {{:gref Z.add }},
32 | {{:gref Z.sub }},
33 | {{:gref Z.mul }},
34 | {{:gref Z.eqb }},
35 | {{:gref Z.compare }},
36 | {{:gref Z.of_nat }},
37 | {{:gref Z.pos_sub }},
38 | {{:gref Z.succ_double }},
39 | {{:gref Z.pred_double }},
40 | {{:gref Z.double }},
41 | {{:gref Pos.succ }},
42 | {{:gref Pos.add }},
43 | {{:gref Pos.sub }},
44 | {{:gref Pos.mul }},
45 | {{:gref Pos.compare }},
46 | {{:gref Pos.compare_cont }},
47 | {{:gref Pos.pred_double }},
48 | {{:gref Pos.of_succ_nat }},
49 | {{:gref negb }}
50 | ].
51 |
--------------------------------------------------------------------------------
/real_simplifier/elpi/theories/RealSimplify.v:
--------------------------------------------------------------------------------
1 | From RealSimplify Require Export Tactic.
--------------------------------------------------------------------------------
/real_simplifier/elpi/theories/Tactic.v:
--------------------------------------------------------------------------------
1 | From elpi Require Import elpi.
2 |
3 | Require Import Reals.
4 | Require Import ZArith.
5 | Require Import QArith.
6 |
7 | From RealSimplify Require Import Theory.
8 | From RealSimplify.Elpi Extra Dependency "reify.elpi" as reify.
9 | From RealSimplify.Elpi Extra Dependency "symbols.elpi" as symbols.
10 |
11 | Elpi Tactic real_simplify.
12 | Elpi Accumulate File reify.
13 | Elpi Accumulate File symbols.
14 | Elpi Accumulate lp:{{
15 | solve (goal _ _ G _ [] as InitialGoal) NewGoals :- std.do! [
16 | std.assert! (G = {{ @eq R lp:A lp:B }}) "goal is not a real equality",
17 | reify A RA,
18 | Eq = {{ cleanup_simplify_correct lp:RA }},
19 | coq.typecheck Eq EqT ok,
20 | coq.redflags.add
21 | coq.redflags.betaiotazeta
22 | {std.map {symbols} (g\ flag\ sigma C\ g = const C, flag = coq.redflags.const C)}
23 | RedFlags,
24 | @redflags! RedFlags => coq.reduction.cbv.norm EqT RedEqT,
25 | RedEqT = {{ lp:C = _ }},
26 | Proof = {{
27 | (@eq_rect R _ (fun t => @eq R t lp:B) _ _
28 | (eq_sym (lp:Eq : lp:RedEqT))
29 | : lp:C = lp:B)
30 | }},
31 | refine Proof InitialGoal NewGoals
32 | ].
33 | }}.
34 | Elpi Typecheck.
35 |
36 | Tactic Notation "real_simplify" := elpi real_simplify.
37 |
--------------------------------------------------------------------------------
/real_simplifier/elpi/theories/Theory.v:
--------------------------------------------------------------------------------
1 | Require Import Reals.
2 | Require Import ZArith.
3 | Require Import QArith.
4 | Require Import Lra.
5 | Require Import Lia.
6 |
7 | Inductive Expr_UnaryOp :=
8 | | EU_Opp
9 | | EU_Inv.
10 |
11 | Inductive Expr_BinaryOp :=
12 | | EB_Add
13 | | EB_Sub
14 | | EB_Mul
15 | | EB_Div
16 | | EB_Max
17 | | EB_Min.
18 |
19 | Inductive Expr_N : Set :=
20 | | EN_lit : nat -> Expr_N
21 | | EN_gen : nat -> Expr_N.
22 |
23 | Inductive Expr_R : Set :=
24 | | ER_Q : Q -> Expr_R
25 | | ER_R : R -> Expr_R
26 | | ER_Z : Z -> Expr_R
27 | | ER_Unary : Expr_UnaryOp -> Expr_R -> Expr_R
28 | | ER_Binary : Expr_BinaryOp -> Expr_R -> Expr_R -> Expr_R
29 | | ER_Pow : Expr_R -> Expr_N -> Expr_R.
30 |
31 | Definition interpret_N (e : Expr_N) : nat :=
32 | match e with
33 | | EN_lit n => n
34 | | EN_gen n => n
35 | end.
36 |
37 | Definition unary_fun (f : Expr_UnaryOp) : R -> R :=
38 | match f with
39 | | EU_Opp => Ropp
40 | | EU_Inv => Rinv
41 | end.
42 |
43 | Definition binary_fun (f : Expr_BinaryOp) : R -> R -> R :=
44 | match f with
45 | | EB_Add => Rplus
46 | | EB_Sub => Rminus
47 | | EB_Mul => Rmult
48 | | EB_Div => Rdiv
49 | | EB_Max => Rmax
50 | | EB_Min => Rmin
51 | end.
52 |
53 | Definition unary_fun_q (f : Expr_UnaryOp) : Q -> Q :=
54 | match f with
55 | | EU_Opp => Qopp
56 | | EU_Inv => Qinv
57 | end.
58 |
59 | Definition binary_fun_q (f : Expr_BinaryOp) : Q -> Q -> Q :=
60 | match f with
61 | | EB_Add => Qplus
62 | | EB_Sub => Qminus
63 | | EB_Mul => Qmult
64 | | EB_Div => Qdiv
65 | | EB_Max => Qminmax.Qmax
66 | | EB_Min => Qminmax.Qmin
67 | end.
68 |
69 | Definition unary_check_args (f : Expr_UnaryOp) (a : Q) : bool :=
70 | match f with
71 | | EU_Inv => negb ((Qnum a) =? 0)%Z
72 | | _ => true
73 | end.
74 |
75 | Definition binary_check_args (f : Expr_BinaryOp) (a b : Q) : bool :=
76 | match f with
77 | | EB_Div => negb ((Qnum b) =? 0)%Z
78 | | _ => true
79 | end.
80 |
81 | Fixpoint interpret_R (e : Expr_R) : R :=
82 | match e with
83 | | ER_Q q => Q2R q
84 | | ER_R r => r
85 | | ER_Z z => IZR z
86 | | ER_Unary f a => (unary_fun f) (interpret_R a)
87 | | ER_Binary f a b => (binary_fun f) (interpret_R a) (interpret_R b)
88 | | ER_Pow a b => pow (interpret_R a) (interpret_N b)
89 | end.
90 |
91 | Fixpoint simplify (e : Expr_R) : Expr_R :=
92 | match e with
93 | | ER_Q q => ER_Q q
94 | | ER_R r => ER_R r
95 | | ER_Z z => ER_Q (z#1)
96 | | ER_Unary f a =>
97 | let a' := simplify a in
98 | match a' with
99 | | ER_Q aq =>
100 | if unary_check_args f aq
101 | then ER_Q ((unary_fun_q f) aq)
102 | else ER_Unary f a'
103 | | _ => ER_Unary f a'
104 | end
105 | | ER_Binary f a b =>
106 | let a' := simplify a in
107 | let b' := simplify b in
108 | match a', b' with
109 | | ER_Q aq, ER_Q bq =>
110 | if binary_check_args f aq bq
111 | then ER_Q ((binary_fun_q f) aq bq)
112 | else ER_Binary f a' b'
113 | | _, _ => ER_Binary f a' b'
114 | end
115 | | ER_Pow a b =>
116 | let a' := simplify a in
117 | match a', b with
118 | | ER_Q aq, EN_lit bn => ER_Q (Qpower aq (Z.of_nat bn))
119 | | _, _ => ER_Pow a' b
120 | end
121 | end.
122 |
123 | Fixpoint cleanup (e : Expr_R) : Expr_R :=
124 | match e with
125 | | ER_Q (z # 1) => ER_Z z
126 | | ER_Q (n # d) => ER_Binary EB_Div (ER_Z n) (ER_Z (Z.pos d))
127 | | ER_R r => e
128 | | ER_Z z => e
129 | | ER_Unary f a => ER_Unary f (cleanup a)
130 | | ER_Binary f a b => ER_Binary f (cleanup a) (cleanup b)
131 | | ER_Pow a b => ER_Pow (cleanup a) b
132 | end.
133 |
134 | Lemma Q2R_max: forall x y : Q,
135 | Q2R (Qminmax.Qmax x y) = Rmax (Q2R x) (Q2R y).
136 | Proof.
137 | intros x y.
138 | destruct (Qlt_le_dec x y) as [Hlt|Hle].
139 | - rewrite Qminmax.Q.max_r, Rmax_right.
140 | + reflexivity.
141 | + apply Qreals.Qle_Rle, Qlt_le_weak, Hlt.
142 | + apply Qlt_le_weak, Hlt.
143 | - rewrite Qminmax.Q.max_l, Rmax_left.
144 | + reflexivity.
145 | + apply Qreals.Qle_Rle, Hle.
146 | + apply Hle.
147 | Qed.
148 |
149 | Lemma Q2R_min: forall x y : Q,
150 | Q2R (Qminmax.Qmin x y) = Rmin (Q2R x) (Q2R y).
151 | Proof.
152 | intros x y.
153 | destruct (Qlt_le_dec x y) as [Hlt|Hle].
154 | - rewrite Qminmax.Q.min_l, Rmin_left.
155 | + reflexivity.
156 | + apply Qreals.Qle_Rle, Qlt_le_weak, Hlt.
157 | + apply Qlt_le_weak, Hlt.
158 | - rewrite Qminmax.Q.min_r, Rmin_right.
159 | + reflexivity.
160 | + apply Qreals.Qle_Rle, Hle.
161 | + apply Hle.
162 | Qed.
163 |
164 | Lemma Qpower_nat_succ: forall (x : Q) (n : nat),
165 | Qpower x (Z.of_nat (S n)) = x * Qpower x (Z.of_nat n).
166 | Proof.
167 | intros x n.
168 | induction n.
169 | - cbn. unfold Qmult. cbn.
170 | rewrite Z.mul_1_r, Pos.mul_1_r.
171 | destruct x as [n d].
172 | reflexivity.
173 | - cbn.
174 | unfold Qpower_positive.
175 | rewrite pow_pos_succ.
176 | + reflexivity.
177 | + apply Eqsth.
178 | + unfold Proper, respectful; intros; subst; reflexivity.
179 | + intros a b c. unfold Qmult. destruct a, b, c. cbn.
180 | rewrite Z.mul_assoc.
181 | rewrite Pos.mul_assoc.
182 | reflexivity.
183 | Qed.
184 |
185 | Lemma Q2R_pow: forall (x : Q) (n : nat),
186 | (Q2R x ^ n)%R = Q2R (x ^ Z.of_nat n).
187 | Proof.
188 | intros x n.
189 | induction n.
190 | - cbn. unfold Q2R. cbn. lra.
191 | - rewrite Qpower_nat_succ. cbn.
192 | rewrite Qreals.Q2R_mult.
193 | rewrite IHn.
194 | reflexivity.
195 | Qed.
196 |
197 | Lemma cleanup_correct: forall (e : Expr_R),
198 | interpret_R e = interpret_R (cleanup e).
199 | Proof.
200 | intros e; induction e as [q|r|z|f a IHa|f a IHa b IHb|a IHa b].
201 | - (* Q *) cbn.
202 | destruct q as [n d]; destruct d; try reflexivity.
203 | unfold Q2R; cbn; lra.
204 | - (* R *) reflexivity.
205 | - (* Z *) reflexivity.
206 | - (* unary *) cbn; rewrite IHa; reflexivity.
207 | - (* binary *) cbn; rewrite IHa, IHb; reflexivity.
208 | - (* pow *) cbn; rewrite IHa; reflexivity.
209 | Qed.
210 |
211 | Lemma simplify_correct: forall (e : Expr_R),
212 | interpret_R e = interpret_R (simplify e).
213 | Proof.
214 | intros e; induction e as [q|r|z|f a IHa|f a IHa b IHb|a IHa b].
215 | - (* Q *) reflexivity.
216 | - (* R *) reflexivity.
217 | - (* Z *) cbn; unfold Q2R; cbn; lra.
218 | - (* unary *) cbn.
219 | destruct (simplify a); rewrite IHa; try reflexivity.
220 | cbn.
221 | destruct f; cbn.
222 | + rewrite Qreals.Q2R_opp; reflexivity.
223 | + destruct (Z.eqb_spec (Qnum q) 0) as [Heq|Hneq]; cbn.
224 | * reflexivity.
225 | * rewrite Qreals.Q2R_inv.
226 | 1: reflexivity.
227 | unfold Qeq; cbn.
228 | lia.
229 | - (* binary *) cbn.
230 | destruct (simplify a) as [qa| | | | |]; rewrite IHa;
231 | destruct (simplify b) as [qb| | | | |]; rewrite IHb; try reflexivity.
232 | cbn.
233 | destruct f; cbn.
234 | + rewrite Qreals.Q2R_plus; reflexivity.
235 | + rewrite Qreals.Q2R_minus; reflexivity.
236 | + rewrite Qreals.Q2R_mult; reflexivity.
237 | + destruct (Z.eqb_spec (Qnum qb) 0) as [Heq|Hneq]; cbn.
238 | * reflexivity.
239 | * rewrite Qreals.Q2R_div.
240 | 1: reflexivity.
241 | unfold Qeq; cbn.
242 | lia.
243 | + rewrite Q2R_max; reflexivity.
244 | + rewrite Q2R_min; reflexivity.
245 | - (* Pow *) cbn.
246 | destruct (simplify a); rewrite IHa; try reflexivity.
247 | destruct b; cbn; try reflexivity.
248 | apply Q2R_pow.
249 | Qed.
250 |
251 | Lemma cleanup_simplify_correct: forall (e : Expr_R),
252 | interpret_R e = interpret_R (cleanup (simplify e)).
253 | Proof.
254 | intros e.
255 | rewrite <- cleanup_correct.
256 | apply simplify_correct.
257 | Qed.
258 |
--------------------------------------------------------------------------------
/real_simplifier/ltac2_reflexive/theories/RealSimpl.v:
--------------------------------------------------------------------------------
1 | (** * Reflexive tactic to compute real expressions using Q arithmetic *)
2 |
3 | (** This is similar to field_simplify, but
4 | - it handles min and max
5 | - it is easier to extend with additional functions which are computable in Q (like modulus)
6 | - it does not change the term structure except for computation (which can be an advantage or disadvantage)
7 |
8 | A few notes:
9 | - this is a reflexive tactic, which means the relevant context (a ℝ term) is converted to a gallina data structure and the
10 | majority of the work (the simplification and computation) is done by gallina functions
11 | - the advantage of reflexive tactics is that one can prove upfront that the transformations done by gallina functions are correct
12 | - in this case this means the equality of the original and simplified term is proven by application of a generic lemma
13 | - this is much faster than constructing and type checking equality lemmas for individual cases
14 | - reflexive tactics typically have these components:
15 | - a reification tactic which converts the relevant context to an AST in gallina
16 | - an interpretation tactic which converts the AST back to the original term (the inverse of reification)
17 | - some processing function on the AST (written in gallina)
18 | - a proof that the processing function has certain properties (in this case preserves equality in ℝ of the interpretation of the AST)
19 | - a wrapper tactic, which does the reification, posts the above proof, computes in the type of the proof and applies this in some form
20 | - this specific instance of a reflexive tactic takes a short cut:
21 | - terms which are not "understood" by the tactic, say variables or unknown functions are copied literally
22 | - not converting the handled terms fully to an AST type and relying on computation with explicit delta lists is "quick and dirty" but quite effective
23 | - the correctness proofs tend to be substantially more complicated if some form of context management is required
24 | - the down side of this method is that if the user supplied term or context does contain symbols of the domain the tactic computes in (ℚ, ℤ, pos in this case)
25 | the terms can blow up
26 | - to avoid this one option is to copy the Q, Z and Pos functions used and use these copies (assuming that these copies do not occur in the user supplied term)
27 |
28 | Caveats:
29 | - this tactic does ℚ, ℤ and Pos computation on parts of the supplied term
30 | - if the term includes computations with variables in these domains, the term might explode
31 | *)
32 |
33 | Require Import Reals.
34 | Require Import ZArith.
35 | Require Import QArith.
36 | Require Import Lra.
37 | Require Import Lia.
38 | Require Import Ltac2.Ltac2.
39 | Require Import Ltac2.Printf.
40 | Require Import Ltac2.Bool.
41 |
42 | (** ** Definition of the AST *)
43 |
44 | Inductive Expr_UnaryOp :=
45 | | EU_Opp
46 | | EU_Inv
47 | .
48 |
49 | Inductive Expr_BinaryOp :=
50 | | EB_Add
51 | | EB_Sub
52 | | EB_Mul
53 | | EB_Div
54 | | EB_Max
55 | | EB_Min
56 | .
57 |
58 | (** For a nat value we need only two variants - something we can compute and something we can't (like a variable) *)
59 |
60 | Inductive Expr_N : Set :=
61 | | EN_lit : nat -> Expr_N (* This is used for everything we can evaluate *)
62 | | EN_gen : nat -> Expr_N (* This is used for everything we cannot evaluate *)
63 | .
64 |
65 | Inductive Expr_R : Set :=
66 | | ER_Q : Q -> Expr_R (* This is used for everything we can evaluate *)
67 | | ER_R : R -> Expr_R (* This is used for everything we cannot evaluate *)
68 | | ER_Z : Z -> Expr_R (* We could use Q for Z, but then interpretation would not be an inverse of reification *)
69 | | ER_Unary : Expr_UnaryOp -> Expr_R -> Expr_R
70 | | ER_Binary : Expr_BinaryOp -> Expr_R -> Expr_R -> Expr_R
71 | | ER_Pow : Expr_R -> Expr_N -> Expr_R
72 | .
73 |
74 | (** ** Interpretation and simplification functions for ℕ *)
75 |
76 | (** For ℕ this is trivial, since it anyway computes *)
77 |
78 | Definition interpret_N (e : Expr_N) : nat :=
79 | match e with
80 | | EN_lit n => n
81 | | EN_gen n => n
82 | end.
83 |
84 | (** ** Interpretation and simplification functions for ℝ *)
85 |
86 | (** Return the ℝ function corresponding to an unary operator *)
87 |
88 | Definition unary_fun_R (f : Expr_UnaryOp) : R->R :=
89 | match f with
90 | | EU_Opp => Ropp
91 | | EU_Inv => Rinv
92 | end.
93 |
94 | (** Return the ℝ function corresponding to a binary operator *)
95 |
96 | Definition binary_fun_R (f : Expr_BinaryOp) : R->R->R :=
97 | match f with
98 | | EB_Add => Rplus
99 | | EB_Sub => Rminus
100 | | EB_Mul => Rmult
101 | | EB_Div => Rdiv
102 | | EB_Max => Rmax
103 | | EB_Min => Rmin
104 | end.
105 |
106 | (** Return the ℚ function corresponding to an unary operator *)
107 |
108 | Definition unary_fun_Q (f : Expr_UnaryOp) : Q->Q :=
109 | match f with
110 | | EU_Opp => Qopp
111 | | EU_Inv => Qinv
112 | end.
113 |
114 | (** Return the ℚ function corresponding to a binary operator *)
115 |
116 | Definition binary_fun_Q (f : Expr_BinaryOp) : Q->Q->Q :=
117 | match f with
118 | | EB_Add => Qplus
119 | | EB_Sub => Qminus
120 | | EB_Mul => Qmult
121 | | EB_Div => Qdiv
122 | | EB_Max => Qminmax.Qmax
123 | | EB_Min => Qminmax.Qmin
124 | end.
125 |
126 | (** Check if the argument of an unary operator is valid (that is not inversion of 0) *)
127 |
128 | Definition unary_check_args_Q (f : Expr_UnaryOp) (a : Q) : bool :=
129 | match f with
130 | | EU_Inv => negb ((Qnum a) =? 0)%Z
131 | | _ => true
132 | end.
133 |
134 | (** Check if the arguments of a binary operator are valid (that is not division by 0) *)
135 |
136 | Definition binary_check_args_Q (f : Expr_BinaryOp) (a b : Q) : bool :=
137 | match f with
138 | | EB_Div => negb ((Qnum b) =? 0)%Z
139 | | _ => true
140 | end.
141 |
142 | (** Interpret an AST, that is convert it back to a ℝ term - this function and the reification tactic must be inverse *)
143 |
144 | Fixpoint interpret_R (e : Expr_R) : R :=
145 | match e with
146 | | ER_Q q => Q2R q
147 | | ER_R r => r
148 | | ER_Z z => IZR z
149 | | ER_Unary f a => (unary_fun_R f) (interpret_R a)
150 | | ER_Binary f a b => (binary_fun_R f) (interpret_R a) (interpret_R b)
151 | | ER_Pow a b => pow (interpret_R a) (interpret_N b)
152 | end.
153 |
154 | (** Simplify an AST by computation using ℚ arithmetic *)
155 |
156 | Fixpoint simplify_R (e : Expr_R) : Expr_R :=
157 | match e with
158 | | ER_Q q => ER_Q q
159 | | ER_R r => ER_R r
160 | | ER_Z z => ER_Q (z#1)
161 | | ER_Unary f a =>
162 | let a':=simplify_R a in
163 | match a' with
164 | | ER_Q aq =>
165 | if unary_check_args_Q f aq
166 | then ER_Q ((unary_fun_Q f) aq)
167 | else ER_Unary f a'
168 | | _ => ER_Unary f a'
169 | end
170 | | ER_Binary f a b =>
171 | let a':=simplify_R a in
172 | let b':=simplify_R b in
173 | match a', b' with
174 | | ER_Q aq, ER_Q bq =>
175 | if binary_check_args_Q f aq bq
176 | then ER_Q ((binary_fun_Q f) aq bq)
177 | else ER_Binary f a' b'
178 | | _, _ => ER_Binary f a' b'
179 | end
180 | | ER_Pow a b =>
181 | let a':=simplify_R a in
182 | match a', b with
183 | | ER_Q aq, EN_lit bn => ER_Q (Qpower aq (Z.of_nat bn))
184 | | _, _ => ER_Pow a' b
185 | end
186 | end.
187 |
188 | (** Convert resulting "Q2R (n#d)" terms to quotients of integers in ℝ - or an integer if the denominator is 1 *)
189 |
190 | (* ToDo: do reduction of Q as well *)
191 |
192 | Fixpoint cleanup_R (e : Expr_R) : Expr_R :=
193 | match e with
194 | | ER_Q (z # 1) => ER_Z z
195 | | ER_Q (n # d) => ER_Binary EB_Div (ER_Z n) (ER_Z (Z.pos d))
196 | | ER_R r => e
197 | | ER_Z z => e
198 | | ER_Unary f a => ER_Unary f (cleanup_R a)
199 | | ER_Binary f a b => ER_Binary f (cleanup_R a) (cleanup_R b)
200 | | ER_Pow a b => ER_Pow (cleanup_R a) b
201 | end.
202 |
203 | (** ** Proofs that the simplification and cleanup functions are correct *)
204 |
205 | (** *** ℚ ℝ arithmetic equivalence lemmas missing in the standard library *)
206 |
207 | Lemma Q2R_max: forall x y : Q,
208 | Q2R (Qminmax.Qmax x y) = Rmax (Q2R x) (Q2R y).
209 | Proof.
210 | intros x y.
211 | destruct (Qlt_le_dec x y) as [Hlt|Hle].
212 | - rewrite Qminmax.Q.max_r, Rmax_right.
213 | + reflexivity.
214 | + apply Qreals.Qle_Rle, Qlt_le_weak, Hlt.
215 | + apply Qlt_le_weak, Hlt. (* Not sure why lra doesn't work here *)
216 | - rewrite Qminmax.Q.max_l, Rmax_left.
217 | + reflexivity.
218 | + apply Qreals.Qle_Rle, Hle.
219 | + apply Hle.
220 | Qed.
221 |
222 | Lemma Q2R_min: forall x y : Q,
223 | Q2R (Qminmax.Qmin x y) = Rmin (Q2R x) (Q2R y).
224 | Proof.
225 | intros x y.
226 | destruct (Qlt_le_dec x y) as [Hlt|Hle].
227 | - rewrite Qminmax.Q.min_l, Rmin_left.
228 | + reflexivity.
229 | + apply Qreals.Qle_Rle, Qlt_le_weak, Hlt.
230 | + apply Qlt_le_weak, Hlt.
231 | - rewrite Qminmax.Q.min_r, Rmin_right.
232 | + reflexivity.
233 | + apply Qreals.Qle_Rle, Hle.
234 | + apply Hle.
235 | Qed.
236 |
237 | Lemma Qpower_nat_succ: forall (x : Q) (n : nat),
238 | Qpower x (Z.of_nat (S n)) = x * Qpower x (Z.of_nat n).
239 | Proof.
240 | (* This is a bit tedious because most lemmas about Q use == instead of =, but we want = here *)
241 | intros x n.
242 | induction n.
243 | - cbn. unfold Qmult. cbn.
244 | rewrite Z.mul_1_r, Pos.mul_1_r.
245 | destruct x as [n d].
246 | reflexivity.
247 | - cbn.
248 | unfold Qpower_positive.
249 | rewrite pow_pos_succ.
250 | + reflexivity.
251 | + apply Eqsth.
252 | + unfold Proper, respectful; intros; subst; reflexivity.
253 | + intros a b c. unfold Qmult. destruct a, b, c. cbn.
254 | rewrite Z.mul_assoc.
255 | rewrite Pos.mul_assoc.
256 | reflexivity.
257 | Qed.
258 |
259 | Lemma Q2R_pow: forall (x : Q) (n : nat),
260 | (Q2R x ^ n)%R = Q2R (x ^ Z.of_nat n).
261 | Proof.
262 | intros x n.
263 | induction n.
264 | - cbn. unfold Q2R. cbn. ltac1:(lra).
265 | - rewrite Qpower_nat_succ. cbn.
266 | rewrite Qreals.Q2R_mult.
267 | rewrite IHn.
268 | reflexivity.
269 | Qed.
270 |
271 | (** *** Correctness lemmas for ℝ*)
272 |
273 | (** The interpretation of a term before and after cleanup_R is equal in ℝ *)
274 |
275 | Lemma cleanup_R_correct: forall (e : Expr_R),
276 | interpret_R e = interpret_R (cleanup_R e).
277 | Proof.
278 | intros e; induction e as [q|r|z|f a IHa|f a IHa b IHb|a IHa b].
279 | - (* Q *) cbn.
280 | destruct q as [n d]; destruct d; try reflexivity.
281 | unfold Q2R; cbn; ltac1:(lra).
282 | - (* R *) reflexivity.
283 | - (* Z *) reflexivity.
284 | - (* unary *) cbn; rewrite IHa; reflexivity.
285 | - (* binary *) cbn; rewrite IHa, IHb; reflexivity.
286 | - (* pow *) cbn; rewrite IHa; reflexivity.
287 | Qed.
288 |
289 | (** The interpretation of a term before and after simplification is equal in ℝ *)
290 |
291 | Lemma simplify_R_correct: forall (e : Expr_R),
292 | interpret_R e = interpret_R (simplify_R e).
293 | Proof.
294 | intros e; induction e as [q|r|z|f a IHa|f a IHa b IHb|a IHa b].
295 | - (* Q *) reflexivity.
296 | - (* R *) reflexivity.
297 | - (* Z *) cbn; unfold Q2R; cbn; ltac1:(lra).
298 | - (* unary *) cbn.
299 | destruct (simplify_R a); rewrite IHa; try reflexivity.
300 | cbn.
301 | destruct f; cbn.
302 | + rewrite Qreals.Q2R_opp; reflexivity.
303 | + destruct (Z.eqb_spec (Qnum q) 0) as [Heq|Hneq]; cbn.
304 | * reflexivity.
305 | * rewrite Qreals.Q2R_inv.
306 | 1: reflexivity.
307 | unfold Qeq; cbn.
308 | ltac1:(lia).
309 | - (* binary *) cbn.
310 | destruct (simplify_R a) as [qa| | | | |]; rewrite IHa;
311 | destruct (simplify_R b) as [qb| | | | |]; rewrite IHb; try reflexivity.
312 | cbn.
313 | destruct f; cbn.
314 | + rewrite Qreals.Q2R_plus; reflexivity.
315 | + rewrite Qreals.Q2R_minus; reflexivity.
316 | + rewrite Qreals.Q2R_mult; reflexivity.
317 | + destruct (Z.eqb_spec (Qnum qb) 0) as [Heq|Hneq]; cbn.
318 | * reflexivity.
319 | * rewrite Qreals.Q2R_div.
320 | 1: reflexivity.
321 | unfold Qeq; cbn.
322 | ltac1:(lia).
323 | + rewrite Q2R_max; reflexivity.
324 | + rewrite Q2R_min; reflexivity.
325 | - (* Pow *) cbn.
326 | destruct (simplify_R a); rewrite IHa; try reflexivity.
327 | destruct b; cbn; try reflexivity.
328 | apply Q2R_pow.
329 | Qed.
330 |
331 | (** The interpretation of a term before and after simplification and cleanup_R is equal in ℝ *)
332 |
333 | Lemma cleanup_simplify_R_correct: forall (e : Expr_R),
334 | interpret_R e = interpret_R (cleanup_R (simplify_R e)).
335 | Proof.
336 | intros e.
337 | rewrite <- cleanup_R_correct.
338 | apply simplify_R_correct.
339 | Qed.
340 |
341 | (** ** Reification and main tactic *)
342 |
343 | (** *** Debugging *)
344 |
345 | (** These notations are intended to wrap printf and can be defined to t() or () to enable / disable the prints.
346 | dbg1 is for low verbosity output, dbg2 for high verbosity output *)
347 |
348 | Ltac2 Notation "dbg1" t(thunk(tactic(6))) := ().
349 | Ltac2 Notation "dbg2" t(thunk(tactic(6))) := ().
350 |
351 | (** *** Term classification *)
352 |
353 | (** Check if "val" is a literal positive *)
354 |
355 | Ltac2 rec is_literal_positive (val : constr) : bool :=
356 | lazy_match! val with
357 | | xI ?sub => is_literal_positive sub
358 | | xO ?sub => is_literal_positive sub
359 | | xH => true
360 | | _ => false
361 | end.
362 |
363 | (** Check if "val" is a literal ℤ *)
364 |
365 | Ltac2 is_literal_Z (val : constr) : bool :=
366 | lazy_match! val with
367 | | Z0 => true
368 | | Zpos ?pos => is_literal_positive pos
369 | | Zneg ?pos => is_literal_positive pos
370 | | _ => false
371 | end.
372 |
373 | (** Check if "val" is a literal ℚ *)
374 |
375 | Ltac2 is_literal_Q (val : constr) : bool :=
376 | lazy_match! val with
377 | | (?n # ?d)%Q => is_literal_Z n && is_literal_positive d
378 | | _ => false
379 | end.
380 |
381 | (** Check if "val" is a computable nat *)
382 |
383 | Ltac2 rec is_computable_nat (val : constr) : bool :=
384 | lazy_match! val with
385 | | (?a + ?b)%nat => is_computable_nat a && is_computable_nat b
386 | | (?a - ?b)%nat => is_computable_nat a && is_computable_nat b
387 | | (?a * ?b)%nat => is_computable_nat a && is_computable_nat b
388 | | S ?a => is_computable_nat a
389 | | O => true
390 | | _ => false
391 | end.
392 |
393 | (** *** Reification *)
394 |
395 | (** This converts a term in ℝ to a Expr_R AST structure - this tactic and "interpret_R" must be inverse to each other *)
396 |
397 | (** Note: we can't prove that the tactic and "interpret_R" are inverses - if they are not the tactic will fail when trying to unify the term in the goal with the interpreted AST.
398 | Probably if the reification would be done in MetaCoq, one could prove this. *)
399 |
400 | Ltac2 rec reify_Expr_N (e : constr) : constr :=
401 | (dbg2 printf "=> reify_Expr_N %t" e);
402 | let res := lazy_match! e with
403 | | Z.to_nat ?z => if is_literal_Z z then '(EN_lit $e) else '(EN_gen $e)
404 | | ?n => if is_computable_nat n then '(EN_lit $n) else '(EN_gen $n)
405 | end in
406 | (dbg2 printf "<= reify_Expr_N %t" res);
407 | res.
408 |
409 | Ltac2 rec reify_Expr_R (e : constr) : constr :=
410 | (dbg2 printf "=> reify_Expr_R %t" e);
411 | let res := lazy_match! e with
412 | | IZR ?z => if is_literal_Z z then '(ER_Z $z) else '(ER_R $e)
413 | | Q2R ?q => if is_literal_Q q then '(ER_Q $q) else '(ER_R $e)
414 | | (- ?a)%R => let a':=reify_Expr_R a in '(ER_Unary EU_Opp $a')
415 | | (/ ?a)%R => let a':=reify_Expr_R a in '(ER_Unary EU_Inv $a')
416 | | (?a + ?b)%R => let a':=reify_Expr_R a in let b':=reify_Expr_R b in '(ER_Binary EB_Add $a' $b')
417 | | (?a - ?b)%R => let a':=reify_Expr_R a in let b':=reify_Expr_R b in '(ER_Binary EB_Sub $a' $b')
418 | | (?a * ?b)%R => let a':=reify_Expr_R a in let b':=reify_Expr_R b in '(ER_Binary EB_Mul $a' $b')
419 | | (?a / ?b)%R => let a':=reify_Expr_R a in let b':=reify_Expr_R b in '(ER_Binary EB_Div $a' $b')
420 | | Rmax ?a ?b => let a':=reify_Expr_R a in let b':=reify_Expr_R b in '(ER_Binary EB_Max $a' $b')
421 | | Rmin ?a ?b => let a':=reify_Expr_R a in let b':=reify_Expr_R b in '(ER_Binary EB_Min $a' $b')
422 | | (?a ^ ?b)%R => let a':=reify_Expr_R a in let b':=reify_Expr_N b in '(ER_Pow $a' $b')
423 | | ?a => '(ER_R $a)
424 | end in
425 | (dbg2 printf "<= reify_Expr_R %t" res);
426 | res.
427 |
428 | (** *** Main tactic "real_simplify" *)
429 |
430 | (** Two tactics to create a fresh name *)
431 |
432 | Ltac2 rec ident_list_of_hyp_list(l : (ident * constr option * constr) list) :=
433 | match l with
434 | | h :: t => let (id,val,type) := h in id :: ident_list_of_hyp_list t
435 | | [] => []
436 | end.
437 |
438 | Ltac2 fresh_name_not_in_local_hyps (base : ident) :=
439 | let hypids := ident_list_of_hyp_list (Control.hyps ()) in
440 | let fresh_name_handler := Fresh.Free.of_ids hypids in
441 | Fresh.fresh fresh_name_handler base.
442 |
443 | (** Create a reduction flags structure. I am not using the "eval cbv" notation for two reasons:
444 | - I am not sure how to do an "eval in type of hyp" in ltac2
445 | - One might want to compute the delta symbol lists, e.g. all symbols in certain ℤ and ℚ modules *)
446 |
447 | Ltac2 redflags_delta_only_all (syms : Std.reference list) :=
448 | {
449 | (* beta: expand the application of an unfolded functions by substitution *)
450 | Std.rBeta := true;
451 | (* delta: true = expand all but rConst; false = expand only rConst *)
452 | Std.rDelta := false;
453 | (* Note: iota in tactics like cbv is a shorthand for match, fix and cofix *)
454 | (* iota-match: simplify_R matches by choosing a pattern *)
455 | Std.rMatch := true;
456 | (* iota-fix: simplify_R fixpoint expressions by expanding one level *)
457 | Std.rFix := true;
458 | (* iota-cofix: simplify_R cofixpoint expressions by expanding one level *)
459 | Std.rCofix := true;
460 | (* zeta: expand let expressions by substitution *)
461 | Std.rZeta := true;
462 | (* Symbols to expand *)
463 | Std.rConst := syms
464 | }.
465 |
466 | (** This is the list of symbols we need to reduce for interpretation, simplification and computation.
467 | Essentially one creates this list first by adding known symbols, trying it and adding symbols remaining in the term.
468 | Note: the resulting terms can be large and CoqIDE is *much* faster than VSCoq in displaying these.
469 | *)
470 |
471 | Ltac2 interpretation_symbols () : Std.reference list := [
472 | (* Simplification and Interpretation functions *)
473 | reference:(interpret_R); reference:(simplify_R); reference:(cleanup_R); reference:(interpret_N);
474 | reference:(binary_fun_R); reference:(unary_fun_R);
475 | reference:(binary_fun_Q); reference:(unary_fun_Q);
476 | reference:(binary_check_args_Q); reference:(unary_check_args_Q);
477 | (* Q unary, binary, comparison, min/max, internal *)
478 | reference:(Qnum); reference:(Qden); reference:(Qopp); reference:(Qinv);
479 | reference:(Qplus); reference:(Qminus); reference:(Qmult); reference:(Qdiv); reference:(Qpower);
480 | reference:(Qcompare); reference:(CompOpp);
481 | reference:(Qminmax.Qmax); reference:(Qminmax.Qmin); reference:(GenericMinMax.gmax); reference:(GenericMinMax.gmin);
482 | reference:(Qpower_positive); reference:(pow_pos);
483 | (* Z unary, binary, comparison, internal *)
484 | reference:(Z.opp);
485 | reference:(Z.add); reference:(Z.sub); reference:(Z.mul);
486 | reference:(Z.eqb); reference:(Z.compare);
487 | reference:(Z.of_nat); reference:(Z.pos_sub); reference:(Z.succ_double); reference:(Z.pred_double); reference:(Z.double);
488 | (* Pos unary, binary, comparison, internal *)
489 | reference:(Pos.succ);
490 | reference:(Pos.add); reference:(Pos.sub); reference:(Pos.mul);
491 | reference:(Pos.compare); reference:(Pos.compare_cont);
492 | reference:(Pos.pred_double); reference:(Pos.of_succ_nat);
493 | (* Boolean *)
494 | reference:(negb)
495 | ].
496 |
497 | (** The main tactic, which reifies the supplied term, poses "cleanup_simplify_R_correct" applied to the ast, computes in the type of this lemma and rewrites with it *)
498 |
499 | Ltac2 real_simplify (term : constr) : unit :=
500 | (dbg1 printf "real_simplify: term = %t" term);
501 | let ast := reify_Expr_R term in
502 | (dbg1 printf "real_simplify: AST = %t" ast);
503 | let theorem_id:=fresh_name_not_in_local_hyps @__equality in
504 | pose ($theorem_id:=(cleanup_simplify_R_correct $ast));
505 | Std.cbv
506 | (redflags_delta_only_all (interpretation_symbols()))
507 | {Std.on_hyps:=Some [(theorem_id, Std.AllOccurrences, Std.InHypTypeOnly)]; Std.on_concl:=Std.NoOccurrences};
508 | let theorem:=Env.instantiate (Std.VarRef theorem_id) in
509 | rewrite $theorem; clear $theorem_id
510 | .
511 |
512 | (** An Ltac1 wrapper for the tactic *)
513 |
514 | Ltac real_simplify := ltac2:(term |- real_simplify (Option.get (Ltac1.to_constr term))).
515 |
516 | (** ** Tests *)
517 |
518 | Section Test.
519 | Variable x : R.
520 | Open Scope R.
521 |
522 | Goal (x+(Rmax ((3*4+5^2-6)/7) (1/2)) = x + 31/7).
523 | match! goal with [ |- (?a = ?b) ] => real_simplify a end.
524 | reflexivity.
525 | Qed.
526 |
527 | End Test.
528 |
--------------------------------------------------------------------------------
/show/elpi/Makefile:
--------------------------------------------------------------------------------
1 | ifeq "$(COQBIN)" ""
2 | COQBIN=$(dir $(shell which coqtop))/
3 | endif
4 |
5 | %: Makefile.coq
6 |
7 | Makefile.coq: _CoqProject
8 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq
9 |
10 | tests: all
11 | @$(MAKE) -C tests -s clean
12 | @$(MAKE) -C tests -s all
13 |
14 | -include Makefile.coq
15 |
--------------------------------------------------------------------------------
/show/elpi/_CoqProject:
--------------------------------------------------------------------------------
1 | -R theories Show
2 |
3 | theories/Show.v
--------------------------------------------------------------------------------
/show/elpi/theories/Show.v:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/rocq-community/metaprogramming-rosetta-stone/48e92067d6c8170a996c42ed88b221bd7ab4ebbd/show/elpi/theories/Show.v
--------------------------------------------------------------------------------