├── PotentialApplications ├── NewApplications │ ├── newTrees.idr │ └── ordering.idr ├── binary.idr ├── myGraph.idr ├── myGraphs.v ├── regExp │ └── regularLanguages.idr ├── testTree.idr ├── tree.idr └── treeSet.idr ├── Provers ├── commutativeGroup_reduce.idr ├── commutativeGroup_test.idr ├── commutativeMonoid_reduce.idr ├── commutativeMonoid_test.idr ├── dataTypes.idr ├── globalDef.idr ├── group_reduce.idr ├── group_test.idr ├── magma_reduce.idr ├── magma_test.idr ├── mathsResults.idr ├── monoid_reduce.idr ├── monoid_test.idr ├── reflection.idr ├── reflectionForSpecificDataTypes.idr ├── reflectionTools.idr ├── ring_reduce.idr ├── ring_test.idr ├── semiGroup_reduce.idr ├── semiGroup_test.idr └── tools.idr ├── README ├── algebraicProvers.ipkg ├── autoAssoc ├── autoAssocList.idr ├── autoAssocNat.idr └── autoAssoc_tools.idr ├── main.idr ├── others ├── AutoAssoc.idr ├── axiomInduction_byRecursivity.idr ├── beqExpr_proof_by_case.v ├── exampleProofByHand.idr ├── howToDoPerfsTest.txt ├── problemsWithIdris │ ├── testReflection2Funs.idr │ ├── testStackProof.idr │ └── unfold_type.idr ├── sortedList.idr ├── test_Ltac_Coq.v └── traditionalApproach.v └── paper-CICM-17 ├── aliascnt.sty ├── content ├── 1-introduction.tex ├── 2-basicTactic.tex ├── 3-hierarchyOfTactics.tex ├── 4-relatedWork.tex └── 5-conclusions.tex ├── dtp.bib ├── llncs.cls ├── llncsdoc.sty ├── paper.tex ├── rebuttal-proposed.txt ├── rebuttal_reviews_Brady_Slama_CICM17.txt ├── remreset.sty ├── splncs03.bst └── sprmindx.sty /PotentialApplications/NewApplications/newTrees.idr: -------------------------------------------------------------------------------- 1 | import ordering 2 | 3 | 4 | left : {A:Type} -> {B:Type} -> (A,B) -> A 5 | left (x,y) = x 6 | 7 | right : {A:Type} -> {B:Type} -> (A,B) -> B 8 | right (x,y) = y 9 | 10 | leftDep : {A:Type} -> {B:A->Type} -> (x : DPair A B) -> A 11 | leftDep (a ** b) = a 12 | 13 | rightDep : {A:Type} -> {B:A->Type} -> (x : DPair A B) -> B (leftDep x) 14 | rightDep (a ** b) = b 15 | 16 | 17 | 18 | -- Just an insertion sort for lists 19 | insert : {A:Type} -> (Aord : PartialOrder A) -> (elem:A) -> List A -> List A 20 | insert Aord elem [] = [elem] 21 | insert Aord elem (h::t) with (lowerEqDec Aord elem h) 22 | insert Aord elem (h::t) | (Yes elemIsLower) = elem::(h::t) 23 | insert Aord elem (h::t) | (No elemIsStrictlyGreater) = h::(insert Aord elem t) 24 | 25 | sortList : {A:Type} -> (Aord : PartialOrder A) -> List A -> List A 26 | sortList Aord [] = [] 27 | sortList Aord (h::t) = insert Aord h (sortList Aord t) 28 | 29 | -- Now, we could prove that sortList is indeed a sorting function... 30 | -- To do 31 | -- [...] 32 | 33 | 34 | -- More interesting : Now, we want to talk about binary trees : 35 | 36 | data Tree : (A:Type) -> List A -> Type where 37 | EmptyTree : {A:Type} -> Tree A [] 38 | Node : {A:Type} -> {l1:List A} -> {l2:List A} -> (left:Tree A l1) -> (elem:A) -> (right:Tree A l2) -> Tree A (l1++[elem]++l2) 39 | 40 | 41 | -- We want to define the function that builds a sorted binary tree from an (unsorted) list of elements 42 | -- And we want to show in the type that the produced tree reflects the same elements, but sorted 43 | -- And since we trust the sort function on list (because it has already been proven to be correct), we can therefore 44 | -- trust this function that operated on trees. 45 | 46 | -- That should brings problems as the index being built is certainly not the same as the one claimed in the type 47 | -- -> we want to use one of our automatic prover for that 48 | 49 | 50 | 51 | 52 | insertTree : {A:Type} -> (Aord : PartialOrder A) -> (x:A) -> {lindex:List A} -> Tree A lindex -> Tree A (insert Aord x lindex) 53 | insertTree {A=A} _ x EmptyTree = Node EmptyTree x EmptyTree 54 | insertTree {A=A} Aord x (Node left elem right) with (lowerEqDec Aord x elem) 55 | insertTree {A=A} Aord x (Node left elem right) | (Yes xIsSmallerOrEqual) ?= Node (insertTree Aord x left) elem right 56 | insertTree {A=A} Aord x (Node left elem right) | (No xIsBigger) ?= Node left elem (insertTree Aord x right) 57 | 58 | 59 | 60 | aux_buildSortedTree : {A:Type} -> (Aord : PartialOrder A) -> {lindex:List A} -> (tree:Tree A (sortList Aord lindex)) -> (l:List A) -> Tree A (sortList Aord (l++lindex)) 61 | aux_buildSortedTree Aord tree [] = tree 62 | aux_buildSortedTree Aord tree (head::tail) ?= aux_buildSortedTree Aord (insertTree Aord head tree) tail 63 | 64 | 65 | -- buildSortedTree : {A:Type} -> (Aord : PartialOrder A) -> (l:List A) -> Tree A (sortList Aord l) 66 | -- buildSortedTree {A=A} Aord [] = EmptyTree 67 | -- buildSortedTree {A=A} Aord (h::t) ?= aux_buildSortedTree Aord (Node EmptyTree h EmptyTree) t 68 | -------------------------------------------------------------------------------- /PotentialApplications/NewApplications/ordering.idr: -------------------------------------------------------------------------------- 1 | -- Franck Slama 2 | -- University of St Andrews 3 | -- ------------------------------ 4 | 5 | module ordering 6 | 7 | 8 | %access public export 9 | 10 | %default total 11 | 12 | 13 | -- This is code that conforms to the description in the article 14 | -- "Automatic predicate testing in formal certification" (accepted for Tests and Proofs 2016) 15 | -- by the author 16 | 17 | 18 | data Or : (A:Type) -> (B:Type) -> Type where 19 | Or_introL : {A:Type} -> {B:Type} -> (a:A) -> Or A B 20 | Or_introR : {A:Type} -> {B:Type} -> (b:B) -> Or A B 21 | 22 | 23 | ------------------------------------------------------ 24 | -- SET (SETOID) -- 25 | ------------------------------------------------------ 26 | 27 | infixl 5 ~= 28 | interface Set c where 29 | -- We just require an equivalence relation over the elements of the "set" 30 | -- which means that two elements are EQuivalent 31 | -- And a way to produce a proof when it holds (it is semi-decidable, or weakly-decidable, call it as you want) 32 | -- I could call this structure "Setoid" instead of "Set" 33 | (~=) : (x:c) -> (y:c) -> Type -- The (undecidable) relation 34 | eqDec : (x:c) -> (y:c) -> Dec(x~=y) 35 | 36 | -- The binary relation should be an equivalence relation, ie 37 | -- Reflexive, 38 | eq_refl : (x:c) -> x~=x 39 | -- Symmetric, 40 | eq_sym : {x:c} -> {y:c} -> (x~=y) -> (y~=x) 41 | -- And transitive 42 | eq_trans : {x:c} -> {y:c} -> {z:c} -> (x~=y) -> (y~=z) -> (x~=z) 43 | 44 | 45 | ------------------------------------------------------ 46 | -- PARTIAL STRICT ORDER -- 47 | ------------------------------------------------------ 48 | 49 | -- Should be in fact called "PartialyStriclyOrdered c" for saying that 50 | -- c is partially strictly ordered, but that's a bit too verbose, 51 | -- so we just say "PartialStrictOrder c" 52 | 53 | interface Set c => PartialStrictOrder c where 54 | (<<) : (x:c) -> (y:c) -> Type -- The (undecidable) relation 55 | 56 | -- The new (<<) operation preserves the equivalence define at the Set level 57 | lower_compat_equivalence_L : {x:c} -> {y:c} -> {y':c} -> (x << y) -> (x ~= x') -> (x' << y) 58 | lower_compat_equivalence_R : {x:c} -> {y:c} -> {y':c} -> (x << y) -> (y ~= y') -> (x << y') 59 | 60 | lowerDec : (x:c) -> (y:c) -> Dec(x< {y:c} -> (x< (y< (x~=y) 63 | lower_trans : {x:c} -> {y:c} -> {z:c} -> (x< (y< (x< PartialOrder c where 76 | 77 | -- Nothing here ! 78 | 79 | -- I can't put the operations (<~=), the functions lowerEqDec and the properties about (<~=) here 80 | -- because there's no way to say "final" for a definitions in Idris, that would mean that the operation can't 81 | -- be redefined by the user. 82 | -- Without being sure that the user can't redefine it, we can't unfold the definition of (<~=) 83 | -- Question to idris designers : Does Idris should have the keyword "final" for allowing this kind of constructions ? 84 | 85 | 86 | -- Lower or equivalent operation 87 | (<~=) : {c:Type} -> {p:PartialOrder c} -> (x:c) -> (y:c) -> Type -- The (undecidable) relation 88 | (<~=) x y = Or (x< {p:PartialOrder c} -> (x:c) -> (y:c) -> (p1 : (x< Void) -> (p2 : (x~=y) -> Void) -> (pFalse : (<~=) {p=p} x y) -> Void 92 | not_lowerEq_lemma1 x y p1 p2 (Or_introL prStrictlyLower) = p1 prStrictlyLower 93 | not_lowerEq_lemma1 x y p1 p2 (Or_introR prEqual) = p2 prEqual 94 | 95 | 96 | -- There's only one way to define the function which decides the lower or equal operation 97 | lowerEqDec : {c:Type} -> (p:PartialOrder c) -> (x:c) -> (y:c) -> Dec((<~=) {p=p} x y) -- FIX IDRIS : If I write (x <~= y) then in the file proofAndTests.idr I get a weird and broken error message which says "can't infer argument 'p''... 98 | lowerEqDec {c} p x y with (lowerDec {c=c} x y) 99 | lowerEqDec {c} p x y | (Yes prStrictlyLower) = Yes (Or_introL prStrictlyLower) 100 | lowerEqDec {c} p x y | (No prNotStrictlyLower) with (eqDec x y) 101 | lowerEqDec {c} p x y | (No prNotStrictlyLower) | (Yes prEqual) = Yes (Or_introR prEqual) 102 | lowerEqDec {c} p x y | (No prNotStrictlyLower) | (No prNotEqual) = No ?MlowerEqDec_1 103 | 104 | 105 | -- I'm forced to make the (Partialorder c) an explicit argument, otherwise we end up with a non-sense error message 106 | lowerEq_refl : {c:Type} -> (p:PartialOrder c) -> (x:c) -> ((<~=) {p=p} x x) 107 | lowerEq_refl p x = Or_introR (eq_refl x) 108 | 109 | 110 | -- FIX IDRIS : I can't simply state "x <~= y" because otherwise Idris introduces many instances instead of using the available 'p' 111 | lowerEq_antisym : {c:Type} -> (p:PartialOrder c) -> {x:c} -> {y:c} -> ((<~=) {p=p} x y) -> ((<~=) {p=p} y x) -> (x ~= y) 112 | -- FIX IDRIS : I have to do these proofs in proof mode, otherwise Idris can't find automatically the corresponding instances... 113 | lowerEq_antisym p (Or_introL x_lower_y) (Or_introL y_lower_x) = ?MlowerEq_antisym_1 -- lower_antisym x_lower_y y_lower_x 114 | lowerEq_antisym p (Or_introL x_lower_y) (Or_introR y_equals_x) = ?MlowerEq_antisym_2 -- eq_sym y_equals_x 115 | lowerEq_antisym p (Or_introR x_equals_y) _ = x_equals_y 116 | 117 | 118 | -- FIX IDRIS : I can't simply state "x <~= y" because otherwise Idris introduces many instances instead of using the available 'p' 119 | lowerEq_trans : {c:Type} -> (p:PartialOrder c) -> {x:c} -> {y:c} -> {z:c} -> ((<~=) {p=p} x y) -> ((<~=) {p=p} y z) -> ((<~=) {p=p} x z) 120 | -- FIX IDRIS : I need to provide {c=c} because Idris can't find automatically the right instances. And I can't give them. But just giving {c=c} helps... 121 | lowerEq_trans {c=c} p (Or_introL x_lower_y) (Or_introL y_lower_z) = Or_introL (lower_trans {c=c} x_lower_y y_lower_z) 122 | lowerEq_trans {c=c} p (Or_introL x_lower_y) (Or_introR y_equals_z) = Or_introL (lower_compat_equivalence_R {c=c} x_lower_y y_equals_z) 123 | lowerEq_trans {c=c} p (Or_introR x_equals_y) (Or_introL y_lower_z) = ?MlowerEq_trans_1 -- Or_introL (lower_compat_equivalence_L {c=c} y_lower_z (eq_sym {c=c} x_equals_y)) 124 | lowerEq_trans {c=c} p (Or_introR x_equals_y) (Or_introR y_equals_z) = Or_introR (eq_trans {c=c} x_equals_y y_equals_z) 125 | 126 | 127 | ------------------------------------------------------ 128 | -- TOTAL ORDER -- 129 | ------------------------------------------------------ 130 | 131 | -- Should be in fact called "TotalyOrdered c" for saying that 132 | -- c is totaly ordered, but that's a bit too verbose, 133 | -- so we just say "TotalOrder c" 134 | 135 | interface PartialOrder c => TotalOrder c where 136 | 137 | -- We can now compare any two elements with (<~=) 138 | 139 | -- FIX IDRIS : I can't simply state "x <~= y" because otherwise it doesn't typecheck... 140 | lowerE_undec_total : (x:c) -> (y:c) -> Or ((<~=) {p=p} x y) ((<~=) {p=p} y x) 141 | 142 | 143 | 144 | 145 | ---------- Proofs ---------- 146 | ordering.MlowerEqDec_1 = proof 147 | intros 148 | exact (not_lowerEq_lemma1 x y prNotStrictlyLower prNotEqual __pi_arg) 149 | 150 | ordering.MlowerEq_antisym_2 = proof 151 | intros 152 | mrefine eq_sym 153 | exact y_equals_x 154 | 155 | ordering.MlowerEq_antisym_1 = proof 156 | intros 157 | mrefine lower_antisym 158 | exact x_lower_y 159 | exact y_lower_x 160 | 161 | ordering.MlowerEq_trans_1 = proof 162 | intros 163 | mrefine Or_introL 164 | mrefine lower_compat_equivalence_L 165 | exact y 166 | exact z 167 | exact y_lower_z 168 | mrefine eq_sym 169 | exact x_equals_y 170 | 171 | 172 | -------------------------------------------------------------------------------- /PotentialApplications/myGraph.idr: -------------------------------------------------------------------------------- 1 | -- myGraph.idr 2 | -- Author : Franck Slama 3 | 4 | 5 | import Data.Fin 6 | 7 | 8 | data Graph : Nat -> Nat -> Type where 9 | -- Add n nodes at once 10 | addNodes : (nbNodes : Nat) -> Graph nbNodes Z 11 | -- Add an edge between two already existing nodes 12 | -- The new edge connects the nodes i and j 13 | addEdge : {nbNodes : Nat} -> {nbEdges : Nat} -> (i,j : Fin nbNodes) -> (Graph nbNodes nbEdges) -> Graph nbNodes (S nbEdges) 14 | 15 | 16 | -- --------------------- 17 | -- Utility functions --- 18 | -- --------------------- 19 | 20 | -- Without using the index (recomputes n using the structure of the graph) 21 | getNbNodes : {n, e : Nat} -> (g : Graph n e) -> Nat 22 | getNbNodes (addNodes n) = n 23 | getNbNodes (addEdge _ _ g') = getNbNodes g' --continue 24 | 25 | 26 | -- Without using the index (recomputes e using the structure of the graph) 27 | getNbEdges : {n, e : Nat} -> (g : Graph n e) -> Nat 28 | getNbEdges (addNodes n) = Z 29 | getNbEdges (addEdge _ _ g') = S (getNbEdges g') 30 | 31 | 32 | 33 | data edgeExists : {n, e : Nat} -> (Graph n e) -> (Fin n) -> (Fin n) -> Type where 34 | isLastAdded : {n, e : Nat} -> (g : Graph n e) -> (i, j : Fin n) -> edgeExists (addEdge i j g) i j 35 | -- the edge [i,j] was already contained in g, so it is still after having added [i',j'] 36 | wasAddedBefore : {n, e : Nat} -> (g:Graph n e) -> (i, j, i', j' : Fin n) -> edgeExists g i j -> edgeExists (addEdge i' j' g) i j 37 | 38 | 39 | -- Boolean equality on Fin 40 | Fin_beq : {n:Nat} -> Fin n -> Fin n -> Bool 41 | Fin_beq i j with (decEq i j) 42 | | (Yes pyes) = True 43 | | (No pno) = False 44 | 45 | 46 | edgeExists_dec : {n, e : Nat} -> (Graph n e) -> (i, j : Fin n) -> Bool 47 | -- There is no edges, so the answer is no... 48 | edgeExists_dec (addNodes _) _ _ = False 49 | -- Here we have to check 50 | edgeExists_dec (addEdge i' j' g') i j = ((Fin_beq i i') && (Fin_beq j j') 51 | || 52 | (Fin_beq i j') && (Fin_beq j i')) 53 | || 54 | (edgeExists_dec g' i j) 55 | 56 | 57 | -- To do : proof that edgeExists_dec effectively decides the predicate edgeExists 58 | 59 | 60 | -- First, naive, equivalence relation 61 | -- Note : good thing : thanks to dependent types, we don't have to state that they should have the same number of nodes and edges 62 | graphEq_1 : {n, e : Nat} -> (g1, g2 : Graph n e) -> Type 63 | -- Note : uggly because Idris is not intended to write logical statments (we don't have the equivalence symbol <->, and other useful things) 64 | graphEq_1 g1 g2 = ((i, j : Fin n) -> (edgeExists g1 i j) -> (edgeExists g2 i j) 65 | , 66 | (i, j : Fin n) -> (edgeExists g2 i j) -> (edgeExists g1 i j)) 67 | 68 | 69 | 70 | -- To do : writing a decision procedure for graphEq_1 and proving that it effectively decides this predicate 71 | 72 | 73 | 74 | 75 | 76 | 77 | -------------------------------------------------------------------------------- /PotentialApplications/myGraphs.v: -------------------------------------------------------------------------------- 1 | Require Import Decidable. 2 | Local Open Scope nat_scope. 3 | Require Import Arith. 4 | 5 | Inductive Fin : nat -> Type := 6 | | fZ : forall n, Fin (S n) 7 | | fS : forall n, Fin n -> Fin (S n). 8 | 9 | Fixpoint fin_dec {n:nat} (i:Fin n) (j:Fin n) {struct n} : bool := 10 | match n with 11 | |O => true 12 | |S pn => match (i,j) with 13 | |(fZ k, fZ k') => beq_nat k k' 14 | |(fS _ i', fS _ j') => fin_dec i' j' 15 | | _ => false 16 | end 17 | end. 18 | 19 | 20 | Inductive Graph : nat -> nat -> Type := 21 | (* Add n nodes at once *) 22 | | addNodes : forall (nbNodes:nat), Graph nbNodes O 23 | (* Add an edge between two already existing nodes *) 24 | (* The new edge connects the nodes i and j *) 25 | | addEdge : forall {nbNodes:nat} {nbEdges} (i j : Fin nbNodes), 26 | (Graph nbNodes nbEdges) -> Graph nbNodes (S nbEdges). 27 | 28 | 29 | (* Utility functions *) 30 | 31 | (* Without using the index (recomputes n using the structure of the graph) *) 32 | Fixpoint getNbNodes {n e : nat} (g : Graph n e) : nat := 33 | match g with 34 | |addNodes nbNodes => nbNodes 35 | |addEdge _ _ _ _ g' => getNbNodes g' (*continue *) 36 | end. 37 | 38 | (* Without using the index (recomputes e using the structure of the graph) *) 39 | Fixpoint getNbEdges {n e : nat} (g : Graph n e) : nat := 40 | match g with 41 | |addNodes n => O 42 | |addEdge _ _ _ _ g' => S (getNbEdges g') 43 | end. 44 | 45 | 46 | 47 | (* First equivalence relation *) 48 | 49 | Inductive edgeExists : forall {n e : nat}, (Graph n e) -> (Fin n) -> (Fin n) -> Prop := 50 | |isLastAdded : forall {n e : nat} (g : Graph n e) (i j : Fin n), edgeExists (addEdge i j g) i j 51 | (* the edge [i,j] was already contained in g, so it is still after having added [i',j'] *) 52 | |wasAddedBefore : forall {n e : nat} (g:Graph n e) (i j i' j' : Fin n), edgeExists g i j -> edgeExists (addEdge i' j' g) i j. 53 | 54 | 55 | Definition edgeExists_dec {n e : nat} (g : Graph n e) (i j : Fin n) : bool := 56 | match g with 57 | |addNodes _ => false 58 | |addEdge _ _ i' j' g' => 59 | end. 60 | 61 | 62 | 63 | (* First, naive, equivalence relation *) 64 | (* Note : good thing : thanks to dependent types, we don't have to state that they should have the same number of nodes and edges *) 65 | Definition graphEq_1 {n e : nat} (g1 g2 : Graph n e) : Prop := 66 | forall (i j : Fin n), edgeExists g1 i j <-> edgeExists g2 i j. 67 | 68 | 69 | 70 | 71 | 72 | 73 | -------------------------------------------------------------------------------- /PotentialApplications/testTree.idr: -------------------------------------------------------------------------------- 1 | import Data.Vect 2 | 3 | {- 4 | -- Just a try 5 | myT : (A:Type) -> Type 6 | myT A = (Sigma Nat (\n => Vect n Bool)) -- Weird, I can't use the ** syntax here. It's probably because of the fact that ** is used both at the type level and at the term level (unlike Coq) 7 | -} 8 | 9 | total 10 | leftDep : {T:Type} -> {P:T->Type} -> (Sigma T P) -> T 11 | leftDep (MkSigma t p) = t 12 | 13 | total 14 | allLeft : {T:Type} -> {P:T->Type} -> (ls : List(Sigma T P)) -> List T 15 | allLeft [] = [] 16 | allLeft (h::t) = (leftDep h) :: (allLeft t) 17 | 18 | total 19 | concatList : {A:Type} -> (List (List A)) -> List A 20 | concatList [] = [] 21 | concatList (h::t) = h ++ (concatList t) 22 | 23 | 24 | concatListDepPair : {A:Type} -> {T:(List A)->Type} -> (ls : List(Sigma (List A) (\l => T l))) -> List A 25 | concatListDepPair ls = let allTs = allLeft ls in concatList allTs 26 | 27 | 28 | -- Tree (with elements of type A in the leaves) and elements of type B in the other nodes 29 | data Tree : (A:Type) -> Type -> List A -> Type where 30 | Leaf : {A:Type} -> {B:Type} -> (a:A) -> Tree A B [a] 31 | Node : {A:Type} -> {B:Type} -> (b:B) -> (ls:List (Sigma (List A) (\l => Tree A B l))) -> Tree A B (concatListDepPair ls) 32 | 33 | 34 | 35 | UnionOfTree : {A:Type} -> {B:Type} -> {ls1:List (Sigma (List A) (\l => Tree A B l))} -> {ls2:List (Sigma (List A) (\l => Tree A B l))} -> (Tree A B (concatListDepPair ls1)) -> (b:B) -> (Tree A B (concatListDepPair ls2)) -> Tree A B (concatListDepPair (ls1++ls2)) 36 | UnionOfTree t1 b t2 ?= Node b [(_ ** t1), (_ ** t2)] 37 | 38 | 39 | 40 | 41 | 42 | -------------------------------------------------------------------------------- /PotentialApplications/tree.idr: -------------------------------------------------------------------------------- 1 | -- Tree (with elements of type A in the leaves) and elements of type B in the other nodes 2 | data Tree : (A:Type) -> Type -> List A -> Type where 3 | Leaf : {A:Type} -> {B:Type} -> (a:A) -> Tree A B [a] 4 | Node : {A:Type} -> {B:Type} -> {l1:List A} -> {l2:List A} -> (left:Tree A B l1) -> (elem:B) -> (right:Tree A B l2) -> Tree A B (l1++l2) 5 | 6 | leavesOfTree : {A:Type} -> {B:Type} -> {l:List A} -> Tree A B l -> List A 7 | leavesOfTree (Leaf a) = [a] 8 | leavesOfTree (Node left elem right) = (leavesOfTree left) ++ (leavesOfTree right) 9 | 10 | leavesOfTree_correct : {A:Type} -> {B:Type} -> {l:List A} -> (t:Tree A B l) -> leavesOfTree t = l 11 | leavesOfTree_correct (Leaf a) = Refl 12 | leavesOfTree_correct (Node {l1=l1} {l2=l2} left elem right) = 13 | let left_correct : (leavesOfTree left = l1) = leavesOfTree_correct left in 14 | let right_correct : (leavesOfTree right = l2) = leavesOfTree_correct right in 15 | rewrite left_correct in (rewrite right_correct in Refl) 16 | 17 | -- not a good example as we have no poblem with the index... ;) 18 | linearise : {A:Type} -> {B:Type} -> {l:List A} -> (t:Tree A B l) -> Tree A Unit l 19 | linearise {l=l} t = -- let leaves = leavesOfTree t in aux_linearise leaves where 20 | let leaves = l in aux_linearise leaves where 21 | aux_linearise : {A:Type} -> (l:List A) -> Tree A Unit l 22 | aux_linearise [h] = Leaf h 23 | aux_linearise (head::tail) = Node (Leaf head) MkUnit (aux_linearise tail) 24 | 25 | {-} 26 | linearise2 : {A:Type} -> {B:Type} -> {l:List A} -> (t:Tree A B l) -> Tree A Unit l 27 | linearise2 (Leaf a) = Leaf a 28 | linearise2 (Node (Leaf v0) A (Leaf v1)) = (Node (Leaf v0) A (Leaf v1)) 29 | linearise2 (Node (Node (Leaf v0) B (Leaf v1)) 30 | A 31 | (Node (Leaf v2) C (Leaf v3)) 32 | ) = 33 | -} 34 | 35 | addAtTheBottomRightCorner : {A:Type} -> {B:Type} -> {l:List A} -> (t:Tree A B l) -> (x:A) -> (inhabitantOfB : B) -> Tree A B (l++[x]) 36 | addAtTheBottomRightCorner (Leaf a) x inhabitantOfB = Node (Leaf a) inhabitantOfB (Leaf x) 37 | addAtTheBottomRightCorner (Node left elem right) x inhabitantOfB ?= Node left elem (addAtTheBottomRightCorner right x inhabitantOfB) 38 | 39 | 40 | 41 | 42 | 43 | 44 | -- Not a good example as we have no problem with the index... ;) 45 | eraseTheBs : {A:Type} -> {B:Type} -> {l:List A} -> (Tree A B l) -> (Tree A Unit l) 46 | eraseTheBs (Leaf a) = Leaf a 47 | eraseTheBs (Node left elem right) = Node (eraseTheBs left) MkUnit (eraseTheBs right) 48 | 49 | {- 50 | linearise : {A:Type} -> {B:Type} -> {l:List A} -> (Tree A B l) -> (Tree A B l) 51 | linearise (Leaf a) = Leaf a 52 | linearise (Node left elem right) = 53 | -} 54 | 55 | UnionOfTree : {A:Type} -> {B:Type} -> {l1:List A} -> {l2:List A} -> (Tree A B l1) -> (b:B) -> (Tree A B l2) -> Tree A B (l1++l2) 56 | UnionOfTree t1 b t2 = Node t1 b t2 57 | 58 | 59 | 60 | 61 | 62 | {- 63 | flatten : {A:Type} -> {B:Type} -> (Tree A B) -> List A 64 | flatten (Leaf a) = [a] 65 | flatten (Node left right elem) = (flatten left) ++ (flatten right) 66 | -} 67 | -------------------------------------------------------------------------------- /PotentialApplications/treeSet.idr: -------------------------------------------------------------------------------- 1 | -- "set", with repetition allowed (ie, lists) 2 | data Set : (A:Type) -> Type where 3 | EmptySet : {A:Type} -> Set A 4 | AddElem : {A:Type} -> (a:A) -> (subset:Set A) -> Set A 5 | 6 | data In : {A:Type} -> (a:A) -> (Set A) -> Type where 7 | isHere : {A:Type} -> (a:A) -> (subset:Set A) -> (In a (AddElem a subset)) 8 | isIn : {A:Type} -> (a:A) -> (subset:Set A) -> (In a subset) -> (head:A) -> In a (AddElem head subset) 9 | 10 | 11 | data SubSet : {A:Type} -> (Set A) -> (Set A) -> Type where 12 | EmptySetEquiv : SubSet EmptySet EmptySet 13 | SameFirstElemEquiv : {A:Type} -> (subset1:Set A) -> (subset2:Set A) -> (a:A) -> (subsetEq:SubSet subset1 subset2) -> SubSet (AddElem a subset1) (AddElem a subset2) 14 | DiffFirstElemQuiv : {A:Type} -> (subset1:Set A) -> (subset2:Set A) -> 15 | (a1:A) -> (a2:A) -> (In a1 subset2) -> 16 | (SubSet subset1 (AddElem a2 subset2)) -> 17 | SubSet (AddElem a1 subset1) (AddElem a2 subset2) 18 | 19 | SubSet_refl : {A:Type} -> (s:Set A) -> SubSet s s 20 | SubSet_refl EmptySet = EmptySetEquiv 21 | SubSet_refl (AddElem a subset) = SameFirstElemEquiv subset subset a (SubSet_refl subset) 22 | 23 | (==) : {A:Type} -> (Set A) -> (Set A) -> Type 24 | (==) s1 s2 = (SubSet s1 s2, SubSet s2 s1) 25 | 26 | 27 | -- Equivalent to concat for lists 28 | total 29 | (++) : {A:Type} -> (Set A) -> (Set A) -> Set A 30 | (++) Empty s2 = s2 31 | (++) (AddElem a subset) s2 = AddElem a (subset++s2) 32 | 33 | 34 | left : {A:Type} -> {B:Type} -> (x:(A,B)) -> A 35 | left (a,b) = a 36 | 37 | right : {A:Type} -> {B:Type} -> (x:(A,B)) -> B 38 | right (a,b) = b 39 | 40 | total 41 | -- unionOfSet is commutative : s1++(s2++s3) == (s1++s2)++s3 42 | unionOfSet_assoc : {A:Type} -> (s1:Set A) -> (s2:Set A) -> (s3:Set A) -> ((s1++(s2++s3)) == ((s1++s2)++s3)) 43 | unionOfSet_assoc EmptySet s2 s3 = (SubSet_refl (s2++s3), SubSet_refl (s2++s3)) -- Whaaat ? I could give only s3 ! 44 | --unionOfSet_assoc (AddElem a subset) s2 s3 = (SubSet_refl _, SubSet_refl _) -- Whaaaat ? I should prove that (a::subset)++(s2++s3) == ((a::subset)++s2)++s3 45 | -- which by reduction is a::(subset++(s2++s3)) == (a::(subset++s2))++s3 46 | -- which by a second reduction is a::(subset++(s2++s3)) == a::((subset++s2)++s3) 47 | -- which means that I should REALLY prove that subset++(s2++s3) == (subset++s2)++s3 (by the induction hypothesis) 48 | unionOfSet_assoc (AddElem a subset) s2 s3 = 49 | let inductionHypothesis : (subset++(s2++s3) == (subset++s2)++s3) = unionOfSet_assoc subset s2 s3 in 50 | let x = (SameFirstElemEquiv (subset++(s2++s3)) ((subset++s2)++s3) a (left inductionHypothesis)) in 51 | let y = (SameFirstElemEquiv ((subset++s2)++s3) (subset++(s2++s3)) a (right inductionHypothesis)) in -- Nonsense ! I could use left here 52 | ?MO 53 | 54 | 55 | unionOfSet_comm : {A:Type} -> (s1:Set A) -> (s2:Set A) -> ((s1++s2) == (s2++s1)) 56 | unionOfSet_comm EmptySet (AddElem a2 subset2) = ?MF -- (SubSet_refl _, SubSet_refl _) 57 | 58 | 59 | 60 | data Tree : (A:Type) -> Set A -> Type where 61 | Leaf : {A:Type} -> (a:A) -> Tree A (AddElem a EmptySet) 62 | Node : {A:Type} -> {s1:Set A} -> {s2:Set A} -> (left:Tree A s1) -> (elem:A) -> (right:Tree A s2) -> Tree A (s1++s2) 63 | 64 | 65 | sortTree : {A:Type} -> (Ord A) -> {s1:Set A} -> (Tree A s1) -> (s2 ** (Tree A s2, s1==s2)) 66 | sortTree ord (Leaf a) = (_ ** (Leaf a, (SameFirstElemEquiv EmptySet EmptySet a EmptySetEquiv))) 67 | 68 | 69 | -------------------------------------------------------------------------------- /Provers/commutativeGroup_test.idr: -------------------------------------------------------------------------------- 1 | module Provers.commutativeGroup_test 2 | 3 | import Data.ZZ 4 | import Data.Vect 5 | import Data.Fin 6 | import Provers.globalDef 7 | import Provers.dataTypes 8 | import Provers.tools 9 | import Provers.commutativeGroup_reduce 10 | import Provers.group_test 11 | import Provers.monoid_test 12 | import Provers.semiGroup_test 13 | import Provers.magma_test 14 | 15 | 16 | %access public export 17 | 18 | 19 | implementation dataTypes.CommutativeGroup ZZ where 20 | Plus_comm x y = plusCommutativeZ x y 21 | 22 | 23 | 24 | -- x + (-x) 25 | expA : (x:ZZ) -> ExprCG (%instance) (FakeSetAndMult (commutativeGroup_to_set (%instance))) [x] (x + (- x)) 26 | expA x = PlusCG _ (VarCG _ _ (RealVariable _ _ _ _ FZ)) (NegCG _ (VarCG _ _ (RealVariable _ _ _ _ FZ))) 27 | 28 | 29 | -- (-x) + x 30 | expB : (x:ZZ) -> ExprCG (%instance) (FakeSetAndMult (commutativeGroup_to_set (%instance))) [x] ((-x) + x) 31 | expB x = PlusCG _ (NegCG _ (VarCG _ _ (RealVariable _ _ _ _ FZ))) (VarCG _ _ (RealVariable _ _ _ _ FZ)) 32 | 33 | 34 | -- 0 35 | expC : (x:ZZ) -> ExprCG (%instance) (FakeSetAndMult (commutativeGroup_to_set (%instance))) [x] (Pos 0) 36 | expC x = ConstCG _ _ _ (Pos 0) 37 | 38 | 39 | -- --------------------------------- 40 | -- TEST 1 : Test if x + -x = -x + x 41 | -- --------------------------------- 42 | -- Normalisation of (x+(-x)) that should give Zero, since now we are working on a cgroup 43 | compare_expA_expB : (x:ZZ) -> Maybe (x + (-x) = (-x) + x) 44 | compare_expA_expB x = commutativeGroupDecideEq (%instance) (expA x) (expB x) 45 | 46 | -- Later, we will have a real tactic "CommutativeGroup" which can fail. At this point, we will 47 | -- not have a missing case for "Nothing", which enables now to manipulate some false proof 48 | -- (which causes a crash only when you apply then to a specific value for x) 49 | proof_expA_expB : (x:ZZ) -> (x + (-x) = (-x) + x) 50 | proof_expA_expB x = let (Just ok) = compare_expA_expB x in ok 51 | -- RESULT : Ok, works for all x ! 52 | 53 | -- --------------------------------- 54 | -- TEST 2 : Test if x + -x = 0 55 | -- --------------------------------- 56 | compare_expA_expC : (x:ZZ) -> Maybe (x + (-x) = Pos 0) 57 | compare_expA_expC x = commutativeGroupDecideEq (%instance) (expA x) (expC x) 58 | -- RESULT : Ok, works for all x ! 59 | 60 | -- -------------------------------------------------------------- 61 | -- TEST 3 : Test if ((u + (x + (-y)))) + ((-x + z) + y) = z + u 62 | -- -------------------------------------------------------------- 63 | expD : (x:ZZ) -> (y:ZZ) -> (z:ZZ) -> (u:ZZ) -> ExprCG (%instance) (FakeSetAndMult (commutativeGroup_to_set (%instance))) [x, y, z, u] (((u + (x + (Neg y)))) + ((-x + z) + y)) 64 | expD x y z u = PlusCG _ 65 | (PlusCG _ 66 | (VarCG _ _ (RealVariable _ _ _ _ (FS (FS (FS FZ))))) 67 | (PlusCG _ (VarCG _ _ (RealVariable _ _ _ _ FZ)) (NegCG _ (VarCG _ _ (RealVariable _ _ _ _ (FS FZ)))))) 68 | (PlusCG _ 69 | (PlusCG _ (NegCG _ (VarCG _ _ (RealVariable _ _ _ _ FZ))) (VarCG _ _ (RealVariable _ _ _ _ (FS (FS FZ))))) 70 | (VarCG _ _ (RealVariable _ _ _ _ (FS FZ)))) 71 | 72 | 73 | expE : (x:ZZ) -> (y:ZZ) -> (z:ZZ) -> (u:ZZ) -> ExprCG (%instance) (FakeSetAndMult (commutativeGroup_to_set (%instance))) [x, y, z, u] (z + u) 74 | expE x y z u = PlusCG _ 75 | (VarCG _ _ (RealVariable _ _ _ _ (FS (FS FZ)))) 76 | (VarCG _ _ (RealVariable _ _ _ _ (FS (FS (FS FZ))))) 77 | 78 | 79 | compare_expD_expE : (x:ZZ) -> (y:ZZ) -> (z:ZZ) -> (u:ZZ) -> Maybe (((u + (x + (-y)))) + ((-x + z) + y) = z + u) 80 | compare_expD_expE x y z u = commutativeGroupDecideEq (%instance) (expD x y z u) (expE x y z u) 81 | 82 | -- Later, we will have a real tactic "CommutativeGroup" which can fail. At this point, we will 83 | -- not have a missing case for "Nothing", which enables now to manipulate some false proof 84 | -- (which causes a crash only when you apply then to a specific value for x) 85 | proof_expD_expE : (x:ZZ) -> (y:ZZ) -> (z:ZZ) -> (u:ZZ) -> (((u + (x + (-y)))) + ((-x + z) + y) = z + u) 86 | proof_expD_expE x y z u = let (Just ok) = compare_expD_expE x y z u in ok 87 | -- RESULT : Ok, works for all x ! 88 | 89 | 90 | 91 | 92 | {- 93 | -- Debugging 94 | 95 | expD' : (x:ZZ) -> (y:ZZ) -> (z:ZZ) -> (u:ZZ) -> ExprCG (%instance) [x, y, z, u] (leftDep (commutativeGroupReduce _ (expD x y z u))) 96 | expD' x y z u = left (rightDep (commutativeGroupReduce _ (expD x y z u))) 97 | 98 | -- Use this to test : \x => \y => \z => \u => print_ExprCG show (expD' x y z u) 99 | 100 | expE' : (x:ZZ) -> (y:ZZ) -> (z:ZZ) -> (u:ZZ) -> ExprCG (%instance) [x, y, z, u] (leftDep (commutativeGroupReduce _ (expE x y z u))) 101 | expE' x y z u = left (rightDep (commutativeGroupReduce _ (expE x y z u))) 102 | -} 103 | 104 | 105 | 106 | 107 | 108 | -------------------------------------------------------------------------------- /Provers/commutativeMonoid_reduce.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File commutativeMonoid_reduce.idr 4 | -- Normalize an expression reflecting an element in a commutative monoid 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.commutativeMonoid_reduce 8 | 9 | import Decidable.Equality 10 | import Data.Vect 11 | import Data.Fin 12 | 13 | import Provers.dataTypes 14 | import Provers.tools 15 | import Provers.monoid_reduce 16 | 17 | 18 | %access public export 19 | 20 | -- A usufel lemma for most rewriting in this section 21 | assoc_commute_and_assoc' : {c:Type} -> {p:CommutativeMonoid c} -> (x : c) -> (y : c) -> (z : c) -> (Plus (Plus x y) z ~= Plus (Plus x z) y) 22 | assoc_commute_and_assoc' {c} x y z = let aux : (Plus (Plus x y) z ~= Plus x (Plus y z)) = Plus_assoc {c} x y z in 23 | let aux2 : (Plus x (Plus y z) ~= Plus x (Plus z y)) = ?Massoc_commute_and_assoc'_1 in 24 | let aux3 : (Plus x (Plus z y) ~= Plus (Plus x z) y) = (set_eq_undec_sym {c} (Plus_assoc x z y)) in 25 | ?Massoc_commute_and_assoc'_2 26 | 27 | 28 | -- 2 FUNCTIONS WHICH ADD A TERM TO AN EXPRESSION ALREADY "SORTED" (of the right forme) 29 | -- ----------------------------------------------------------------------------------- 30 | 31 | --%default total 32 | -- The expression is already in the form a + (b + (c + ...)) where a,b,c can only be constants, variables, and negations of constants and variables 33 | putVarOnPlace_cm : {c:Type} -> (p:CommutativeMonoid c) -> (setAndMult:SetWithMult c (commutativeMonoid_to_set p)) -> {g:Vect n c} -> {c1:c} 34 | -> (ExprCM p setAndMult g c1) 35 | -> (i:Fin n) 36 | -> (c2 ** (ExprCM p setAndMult g c2, Plus c1 (index i g) ~= c2)) 37 | putVarOnPlace_cm p setAndMult (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) e) (i) = 38 | if minusOrEqual_Fin i0 i 39 | then let (r_ihn ** (e_ihn, p_ihn))= (putVarOnPlace_cm p setAndMult e i) in 40 | (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) e_ihn, ?MputVarOnPlace_cm_1)) 41 | else (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i)) (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) e), ?MputVarOnPlace_cm_2)) 42 | putVarOnPlace_cm p setAndMult (PlusCM _ (ConstCM _ _ _ c0) e) (i) = 43 | (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i)) (PlusCM _ (ConstCM _ _ _ c0) e), ?MputVarOnPlace_cm_3)) -- the variable becomes the first one, and e is already sorted, there's no need to do a recursive call here ! 44 | -- Basic cases : cases without Plus 45 | putVarOnPlace_cm p setAndMult (ConstCM _ _ _ c0) i = (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i)) (ConstCM _ _ _ c0), ?MputVarOnPlace_cm_4)) 46 | putVarOnPlace_cm {c} p setAndMult (VarCM p _ (RealVariable _ _ _ _ i0)) i = 47 | if minusOrEqual_Fin i0 i then (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) (VarCM p _ (RealVariable _ _ _ _ i)), set_eq_undec_refl {c} _)) 48 | else (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i)) (VarCM p _ (RealVariable _ _ _ _ i0)), ?MputVarOnPlace_cm_5)) 49 | 50 | 51 | putConstantOnPlace_cm : {c:Type} -> (p:CommutativeMonoid c) -> (setAndMult:SetWithMult c (commutativeMonoid_to_set p)) -> {g:Vect n c} -> {c1:c} 52 | -> (ExprCM p setAndMult g c1) -> (constValue:c) 53 | -> (c2 ** (ExprCM p setAndMult g c2, Plus c1 constValue~=c2)) 54 | putConstantOnPlace_cm p setAndMult (PlusCM _ (ConstCM _ _ _ c0) e) constValue = (_ ** (PlusCM _ (ConstCM _ _ _ (Plus c0 constValue)) e, ?MputConstantOnPlace_cm_1)) 55 | putConstantOnPlace_cm p setAndMult (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) e) constValue = 56 | let (r_ihn ** (e_ihn, p_ihn)) = putConstantOnPlace_cm p setAndMult e constValue in 57 | (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) e_ihn, ?MputConstantOnPlace_cm_2)) 58 | -- Basic cases : cases without Plus 59 | putConstantOnPlace_cm {c} p setAndMult (ConstCM _ _ _ c0) constValue = (_ ** (ConstCM _ _ _ (Plus c0 constValue), set_eq_undec_refl {c} _)) 60 | putConstantOnPlace_cm {c} p setAndMult (VarCM p _ (RealVariable _ _ _ _ i0)) constValue = (_ ** (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) (ConstCM _ _ _ constValue), set_eq_undec_refl {c} _)) 61 | 62 | 63 | -- FUNCTION WHICH REORGANIZE AN EXPRESSION 64 | -- ----------------------------------------- 65 | 66 | -- That's the reorganize function for a Commutative Monoid 67 | -- The general pattern is reorganize (Plus term exp) = putTermOnPlace (reorganize exp) term 68 | reorganize_cm : {c:Type} -> (p:CommutativeMonoid c) -> (setAndMult:SetWithMult c (commutativeMonoid_to_set p)) -> {g:Vect n c} -> {c1:c} -> (ExprCM p setAndMult g c1) -> (c2 ** (ExprCM p setAndMult g c2, c1~=c2)) 69 | reorganize_cm p setAndMult (PlusCM _ (VarCM p _ (RealVariable _ _ _ _ i0)) e) = 70 | let (r_ihn ** (e_ihn, p_ihn)) = reorganize_cm p setAndMult e in 71 | let (r_add ** (e_add, p_add)) = putVarOnPlace_cm p setAndMult e_ihn i0 in 72 | (_ ** (e_add, ?Mreorganize_cm_1)) 73 | reorganize_cm p setAndMult (PlusCM _ (ConstCM _ _ _ c0) e) = 74 | let (r_ihn ** (e_ihn, p_ihn)) = reorganize_cm p setAndMult e in 75 | let (r_add ** (e_add, p_add)) = putConstantOnPlace_cm p setAndMult e_ihn c0 in 76 | (_ ** (e_add, ?Mreorganize_cm_2)) 77 | reorganize_cm {c} p setAndMult e = (_ ** (e, set_eq_undec_refl {c} _)) 78 | 79 | 80 | 81 | --WHY DO WE NEED THIS IN THE FUNCTION elimZeroCM JUST UNDER ? THAT'S CRAP ! 82 | getZero_cm : {c:Type} -> (p:dataTypes.CommutativeMonoid c) -> c 83 | getZero_cm p = Zero 84 | 85 | 86 | elimZeroCM : {c:Type} -> (p:dataTypes.CommutativeMonoid c) -> (setAndMult:SetWithMult c (commutativeMonoid_to_set p)) -> {g:Vect n c} -> {c1:c} -> (ExprCM p setAndMult g c1) -> (c2 ** (ExprCM p setAndMult g c2, c1~=c2)) 87 | elimZeroCM p setAndMult (ConstCM p _ _ const) = (_ ** (ConstCM p _ _ const, set_eq_undec_refl const)) 88 | elimZeroCM {c} p setAndMult (PlusCM _ (ConstCM p _ _ const1) (VarCM p _ v)) with (commutativeMonoid_eq_as_elem_of_set p Zero const1) 89 | elimZeroCM {c} p setAndMult (PlusCM _ (ConstCM p _ _ const1) (VarCM p _ v)) | (Just const1_eq_zero) = (_ ** (VarCM p _ v, ?MelimZeroCM1)) 90 | elimZeroCM {c} p setAndMult (PlusCM _ (ConstCM p _ _ const1) (VarCM p _ v)) | _ = (_ ** (PlusCM _ (ConstCM p _ _ const1) (VarCM p _ v), set_eq_undec_refl {c} _)) 91 | elimZeroCM {c} p setAndMult (PlusCM _ (VarCM _ _ v) (ConstCM _ _ _ const2)) with (commutativeMonoid_eq_as_elem_of_set p Zero const2) 92 | elimZeroCM {c} p setAndMult (PlusCM _ (VarCM _ _ v) (ConstCM _ _ _ const2)) | (Just const2_eq_zero) = (_ ** (VarCM _ _ v, ?MelimZeroCM2)) 93 | elimZeroCM {c} p setAndMult (PlusCM _ (VarCM _ _ v) (ConstCM _ _ _ const2)) | _ = (_ ** (PlusCM _ (VarCM _ _ v) (ConstCM _ _ _ const2), set_eq_undec_refl {c} _)) 94 | elimZeroCM p setAndMult (PlusCM _ e1 e2) = 95 | let (r_ih1 ** (e_ih1, p_ih1)) = (elimZeroCM p setAndMult e1) in 96 | let (r_ih2 ** (e_ih2, p_ih2)) = (elimZeroCM p setAndMult e2) in 97 | ((Plus r_ih1 r_ih2) ** (PlusCM _ e_ih1 e_ih2, ?MelimZeroCM3)) 98 | elimZeroCM {c} p setAndMult (VarCM _ _ v) = (_ ** (VarCM _ _ v, set_eq_undec_refl {c} _)) 99 | -- No treatment for Minus since they have already been simplified 100 | -- and no treatment for Neg since we should not find a constant inside a Neg at this point 101 | elimZeroCM {c} p setAndMult e = (_ ** (e, set_eq_undec_refl {c} _)) 102 | 103 | 104 | commutativeMonoidReduce : (p:CommutativeMonoid c) -> {setAndMult:SetWithMult c (commutativeMonoid_to_set p)} -> {g:Vect n c} -> {c1:c} -> (ExprCM p setAndMult g c1) -> (c2 ** (ExprCM p setAndMult g c2, c1~=c2)) 105 | commutativeMonoidReduce p e = 106 | let e_1 = commutativeMonoid_to_monoid e in 107 | let (r_2 ** (e_2, p_2)) = monoidReduce _ (FakeSetAndNeg (commutativeMonoid_to_set p)) e_1 in 108 | let e_3 = monoid_to_commutativeMonoid p e_2 in 109 | let (r_4 ** (e_4, p_4)) = reorganize_cm p _ e_3 in 110 | let (r_5 ** (e_5, p_5)) = elimZeroCM p _ e_4 in 111 | (_ ** (e_5, ?McommutativeMonoidReduce_1)) 112 | 113 | 114 | buildProofCommutativeMonoid : {c:Type} -> {n:Nat} -> (p:CommutativeMonoid c) -> {setAndMult:SetWithMult c (commutativeMonoid_to_set p)} -> {g:Vect n c} -> {x : c} -> {y : c} -> {c1:c} -> {c2:c} -> (ExprCM p setAndMult g c1) -> (ExprCM p setAndMult g c2) -> (x ~= c1) -> (y ~= c2) -> (Maybe (x ~= y)) 115 | buildProofCommutativeMonoid p e1 e2 lp rp with (exprCM_eq p _ _ e1 e2) 116 | buildProofCommutativeMonoid p e1 e2 lp rp | Just e1_equiv_e2 = ?MbuildProofCommutativeMonoid 117 | buildProofCommutativeMonoid p e1 e2 lp rp | Nothing = Nothing 118 | 119 | 120 | commutativeMonoidDecideEq : {c:Type} -> {n:Nat} -> (p:CommutativeMonoid c) -> {setAndMult:SetWithMult c (commutativeMonoid_to_set p)} -> {g:Vect n c} -> {x : c} -> {y : c} -> (ExprCM p setAndMult g x) -> (ExprCM p setAndMult g y) -> (Maybe (x~=y)) 121 | -- e1 is the left side, e2 is the right side 122 | commutativeMonoidDecideEq p e1 e2 = 123 | let (r_e1 ** (e_e1, p_e1)) = commutativeMonoidReduce p e1 in 124 | let (r_e2 ** (e_e2, p_e2)) = commutativeMonoidReduce p e2 in 125 | buildProofCommutativeMonoid p e_e1 e_e2 p_e1 p_e2 126 | 127 | 128 | 129 | ---------- Proofs ---------- 130 | Provers.commutativeMonoid_reduce.Massoc_commute_and_assoc'_1 = proof 131 | intros 132 | mrefine Plus_preserves_equiv 133 | mrefine set_eq_undec_refl 134 | mrefine Plus_comm' 135 | 136 | Provers.commutativeMonoid_reduce.Massoc_commute_and_assoc'_2 = proof 137 | intros 138 | mrefine eq_preserves_eq 139 | exact (Plus x (Plus y z)) 140 | exact (Plus (Plus x z) y) 141 | exact aux 142 | mrefine set_eq_undec_refl 143 | mrefine eq_preserves_eq 144 | exact (Plus x (Plus y z)) 145 | exact (Plus x (Plus z y)) 146 | mrefine set_eq_undec_refl 147 | mrefine set_eq_undec_sym 148 | exact aux2 149 | exact aux3 150 | 151 | Provers.commutativeMonoid_reduce.MputVarOnPlace_cm_1 = proof 152 | intros 153 | mrefine eq_preserves_eq 154 | exact (Plus (Plus (index i0 g) c2) (index i g)) 155 | exact (Plus (index i0 g) (Plus c2 (index i g))) 156 | mrefine set_eq_undec_refl 157 | mrefine Plus_preserves_equiv 158 | mrefine Plus_assoc 159 | mrefine set_eq_undec_refl 160 | mrefine set_eq_undec_sym 161 | exact p_ihn 162 | 163 | Provers.commutativeMonoid_reduce.MputVarOnPlace_cm_2 = proof 164 | intros 165 | mrefine Plus_comm' 166 | 167 | Provers.commutativeMonoid_reduce.MputVarOnPlace_cm_3 = proof 168 | intros 169 | mrefine Plus_comm' 170 | 171 | Provers.commutativeMonoid_reduce.MputVarOnPlace_cm_4 = proof 172 | intros 173 | mrefine Plus_comm' 174 | 175 | Provers.commutativeMonoid_reduce.MputVarOnPlace_cm_5 = proof 176 | intros 177 | mrefine Plus_comm' 178 | 179 | Provers.commutativeMonoid_reduce.MputConstantOnPlace_cm_1 = proof 180 | intros 181 | mrefine assoc_commute_and_assoc' 182 | 183 | Provers.commutativeMonoid_reduce.MputConstantOnPlace_cm_2 = proof 184 | intros 185 | mrefine eq_preserves_eq 186 | exact (Plus (Plus (index i0 g) c2) constValue) 187 | exact (Plus (index i0 g) (Plus c2 constValue)) 188 | mrefine set_eq_undec_refl 189 | mrefine Plus_preserves_equiv 190 | mrefine Plus_assoc 191 | mrefine set_eq_undec_refl 192 | mrefine set_eq_undec_sym 193 | exact p_ihn 194 | 195 | Provers.commutativeMonoid_reduce.Mreorganize_cm_1 = proof 196 | intros 197 | mrefine eq_preserves_eq 198 | exact (Plus (index i0 g) r_ihn) 199 | exact r_add 200 | mrefine Plus_preserves_equiv 201 | mrefine set_eq_undec_refl 202 | mrefine set_eq_undec_sym 203 | mrefine set_eq_undec_refl 204 | exact p_ihn 205 | mrefine set_eq_undec_sym 206 | mrefine eq_preserves_eq 207 | exact (Plus r_ihn (index i0 g)) 208 | exact r_add 209 | mrefine Plus_comm' 210 | mrefine set_eq_undec_refl 211 | exact p_add 212 | 213 | Provers.commutativeMonoid_reduce.Mreorganize_cm_2 = proof 214 | intros 215 | mrefine eq_preserves_eq 216 | exact (Plus c0 r_ihn) 217 | exact r_add 218 | mrefine Plus_preserves_equiv 219 | mrefine set_eq_undec_refl 220 | mrefine eq_preserves_eq 221 | mrefine set_eq_undec_refl 222 | exact p_ihn 223 | exact (Plus r_ihn c0) 224 | exact r_add 225 | mrefine Plus_comm' 226 | mrefine set_eq_undec_refl 227 | exact p_add 228 | 229 | Provers.commutativeMonoid_reduce.MelimZeroCM1 = proof 230 | intros 231 | mrefine eq_preserves_eq 232 | exact (Plus Zero c2) 233 | exact c2 234 | mrefine Plus_preserves_equiv 235 | mrefine set_eq_undec_refl 236 | mrefine Plus_neutral_1 237 | mrefine set_eq_undec_sym 238 | mrefine set_eq_undec_refl 239 | exact const1_eq_zero 240 | 241 | Provers.commutativeMonoid_reduce.MelimZeroCM2 = proof 242 | intros 243 | mrefine eq_preserves_eq 244 | exact (Plus c1 Zero) 245 | exact c1 246 | mrefine Plus_preserves_equiv 247 | mrefine set_eq_undec_refl 248 | mrefine Plus_neutral_2 249 | mrefine set_eq_undec_refl 250 | mrefine set_eq_undec_sym 251 | exact const2_eq_zero 252 | 253 | Provers.commutativeMonoid_reduce.MelimZeroCM3 = proof 254 | intros 255 | mrefine Plus_preserves_equiv 256 | exact p_ih1 257 | exact p_ih2 258 | 259 | Provers.commutativeMonoid_reduce.McommutativeMonoidReduce_1 = proof 260 | intros 261 | mrefine eq_preserves_eq 262 | exact r_2 263 | exact r_5 264 | exact p_2 265 | mrefine set_eq_undec_refl 266 | mrefine eq_preserves_eq 267 | exact r_4 268 | exact r_5 269 | exact p_4 270 | mrefine set_eq_undec_refl 271 | exact p_5 272 | 273 | Provers.commutativeMonoid_reduce.MbuildProofCommutativeMonoid = proof 274 | intros 275 | refine Just 276 | mrefine eq_preserves_eq 277 | exact c1 278 | exact c2 279 | exact lp 280 | exact rp 281 | exact e1_equiv_e2 282 | 283 | 284 | 285 | 286 | 287 | 288 | 289 | 290 | 291 | 292 | 293 | 294 | 295 | -------------------------------------------------------------------------------- /Provers/commutativeMonoid_test.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File commutativeMonoid_test.idr 4 | -- test the normalization for commutative monoid 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.commutativeMonoid_test 8 | 9 | import Data.Vect 10 | import Provers.globalDef 11 | import Provers.dataTypes 12 | import Provers.commutativeMonoid_reduce 13 | import Provers.monoid_test -- For the instance of the level under (Monoid Nat) 14 | 15 | 16 | %access public export 17 | 18 | implementation dataTypes.CommutativeMonoid Nat where 19 | Plus_comm' x y = plusCommutative x y 20 | -------------------------------------------------------------------------------- /Provers/globalDef.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File gloalDef.idr 4 | -- Contains global definitions 5 | ------------------------------------------------------------------- 6 | 7 | module globalDef 8 | 9 | -- Because I don't really like the new Z constructor for the zero of nat. 10 | -- I will see later if I change everything to Z. 11 | O : Nat 12 | O = Z 13 | 14 | -------------------------------------------------------------------------------- /Provers/group_test.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File group_test.idr 4 | -- test the normalization for group 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.group_test 8 | 9 | import Data.ZZ 10 | import Data.Vect 11 | import Data.Fin 12 | 13 | import Provers.globalDef 14 | import Provers.dataTypes 15 | import Provers.tools 16 | import Provers.group_reduce 17 | import Provers.monoid_test 18 | import Provers.semiGroup_test 19 | import Provers.magma_test 20 | 21 | 22 | %access public export 23 | 24 | 25 | implementation dataTypes.Set ZZ where 26 | -- The relation is just the equality 27 | (~=) x y = (x=y) 28 | 29 | set_eq x y with (decEq x y) 30 | set_eq x x | Yes Refl = Just Refl 31 | set_eq x y | _ = Nothing 32 | 33 | set_eq_undec_refl x = Refl 34 | set_eq_undec_sym p = sym p 35 | set_eq_undec_trans p1 p2 = rewrite p1 in rewrite p2 in Refl 36 | 37 | implementation Magma ZZ where 38 | Plus x y = x + y 39 | 40 | Plus_preserves_equiv p1 p2 = ?MPlusZZ_preserves_equiv_1 41 | 42 | implementation SemiGroup ZZ where 43 | Plus_assoc c1 c2 c3 = sym (plusAssociativeZ c1 c2 c3) 44 | 45 | implementation dataTypes.Monoid ZZ where 46 | Zero = Pos Z 47 | 48 | Plus_neutral_1 c = plusZeroLeftNeutralZ c 49 | 50 | Plus_neutral_2 c = plusZeroRightNeutralZ c 51 | 52 | 53 | 54 | plus_Z_simpl : (x:ZZ) -> (y:ZZ) -> (x - y = x + (-y)) 55 | plus_Z_simpl x (Pos Z) = Refl 56 | plus_Z_simpl x (Pos (S y)) = Refl 57 | plus_Z_simpl x (NegS Z) = Refl 58 | plus_Z_simpl x (NegS (S y)) = Refl 59 | 60 | minusNat_Z_Zero : (x:Nat) -> (minusNatZ x x = Pos Z) 61 | minusNat_Z_Zero Z = Refl 62 | minusNat_Z_Zero (S px) = minusNat_Z_Zero px 63 | 64 | plus_inverse : (x:ZZ) -> (x + (- x) = Pos Z, (- x) + x = the ZZ (Pos Z)) 65 | plus_inverse (Pos Z) = (Refl, Refl) 66 | plus_inverse (Pos (S px)) = (minusNat_Z_Zero px, minusNat_Z_Zero px) 67 | plus_inverse (NegS Z) = (Refl, Refl) 68 | plus_inverse (NegS (S py)) = (minusNat_Z_Zero py, minusNat_Z_Zero py) 69 | 70 | 71 | 72 | implementation dataTypes.Group ZZ where 73 | Neg x = -x 74 | Minus x y = x - y 75 | 76 | Neg_preserves_equiv p = ?MNeg_preserves_equiv_1 77 | 78 | Minus_simpl x y = plus_Z_simpl x y --Minus c1 c2 = Plus c1 (Neg c2) --Minus should not be primitive and should be simplifiable 79 | -- The most important stuff for a group is the following : 80 | Plus_inverse x = plus_inverse x --hasSymmetric c (%instance) c1 (Neg c1) -- Every element 'x' has a symmetric which is (Neg x) 81 | 82 | 83 | 84 | 85 | -- ---------------------- 86 | -- TEST 1 THAT SHOULD WORK 87 | -- ---------------------- 88 | termC : (x:ZZ) -> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x] ((2 + (0-2))+x) 89 | termC x = PlusG _ (PlusG _ (ConstG _ _ _ (Pos 2)) 90 | (MinusG _ (ConstG _ _ _ (Pos 0)) (ConstG _ _ _ (Pos 2)))) 91 | (VarG _ _ (RealVariable _ _ _ _ FZ)) 92 | 93 | termD : (x:ZZ) -> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x] x 94 | termD x = VarG _ _ (RealVariable _ _ _ _ FZ) 95 | 96 | 97 | -- Normalisation of ((2 + (0-2))+x) that should give x, since now we are working on a group 98 | compare_termC_termD : (x:ZZ) -> Maybe (((2 + (0-2))+x) = x) 99 | compare_termC_termD x = groupDecideEq (%instance) (termC x) (termD x) 100 | 101 | -- Later, we will have a real tactic "Group" which can fail. At this point, we will 102 | -- not have a missing case for "Nothing", which enables now to manipulate some false proof 103 | -- (which causes a crash only when you apply then to a specific value for x) 104 | proof_termC_termD : (x:ZZ) -> (((2 + (0-2))+x) = x) 105 | proof_termC_termD x = let (Just ok) = compare_termC_termD x in ok 106 | -- RESULT : WORKS FOR ALL X ! 107 | 108 | 109 | termE : (x:ZZ) -> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x] ((3 + (0-2))+x) 110 | termE x = PlusG _ (PlusG _ (ConstG _ _ _ (Pos 3)) 111 | (MinusG _ (ConstG _ _ _ (Pos 0)) (ConstG _ _ _ (Pos 2)))) 112 | (VarG _ _ (RealVariable _ _ _ _ FZ)) 113 | 114 | termF : (x:ZZ) -> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x] (1+x) 115 | termF x = PlusG _ (ConstG _ _ _ (Pos 1)) (VarG _ _ (RealVariable _ _ _ _ FZ)) 116 | 117 | 118 | termG : (x:ZZ) -> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x] x 119 | termG x = VarG _ _ (RealVariable _ _ _ _ FZ) 120 | 121 | 122 | 123 | -- ---------------------- 124 | -- TEST 2 THAT SHOULD WORK 125 | -- ---------------------- 126 | -- Normalisation of ((2 + (0-2))+x) that should give (1+x), since now we are working on a group 127 | compare_termE_termF : (x:ZZ) -> Maybe (((3 + (0-2))+x) = (1+x)) 128 | compare_termE_termF x = groupDecideEq (%instance) (termE x) (termF x) 129 | 130 | -- Later, we will have a real tactic "Group" which can fail... 131 | proof_termE_termF : (x:ZZ) -> (((3 + (0-2))+x) = (1+x)) 132 | proof_termE_termF x = let (Just ok) = compare_termE_termF x in ok 133 | -- RESULT : WORKS FOR ALL X ! 134 | 135 | print_termE_norm : ZZ -> String 136 | print_termE_norm = (\x => print_ExprG show (left (rightDep (groupReduce (%instance) (termE x))))) 137 | 138 | print_termF_norm : ZZ -> String 139 | print_termF_norm = (\x => print_ExprG show (left (rightDep (groupReduce (%instance) (termF x))))) 140 | 141 | -- ---------------------- 142 | -- TEST 3 THAT SHOULD FAIL 143 | -- ---------------------- 144 | -- Normalisation of ((2 + (0-2))+x) that should NOT give x, since now we are working on a group 145 | compare_termE_termG : (x:ZZ) -> Maybe (((3 + (0-2))+x) = x) 146 | compare_termE_termG x = groupDecideEq (%instance) (termE x) (termG x) 147 | 148 | -- Later, we will have a real tactic "Group" which can fail... 149 | proof_termE_termG : (x:ZZ) -> (((3 + (0-2))+x) = x) 150 | proof_termE_termG x = let (Just ok) = compare_termE_termG x in ok 151 | -- RESULT : WORKS FOR ALL X (CAREFUL, it works because it gives Nothing : the two things are NOT equal here!) 152 | 153 | 154 | -- ---------------------- 155 | -- TEST 4 THAT SHOULD WORK 156 | -- ---------------------- 157 | 158 | termH : (x:ZZ) -> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x] ((-2 + (0 + (-(-2)))) + x) 159 | termH x = PlusG _ (PlusG _ (NegG _ (ConstG _ _ _ (Pos 2))) 160 | (PlusG _ (ConstG _ _ _ (Pos 0)) (NegG _ (NegG _ (ConstG _ _ _ (Pos 2)))))) 161 | (VarG _ _ (RealVariable _ _ _ _ FZ)) 162 | 163 | 164 | -- Reminder : termG represents just "x" 165 | 166 | compare_termH_termG : (x:ZZ) -> Maybe (((-2 + (0 + (-(-2)))) + x) = x) 167 | compare_termH_termG x = groupDecideEq (%instance) (termH x) (termG x) 168 | 169 | -- Later, we will have a real tactic "Group" which can fail... 170 | proof_termH_termG : (x:ZZ) -> (((-2 + (0 + (-(-2)))) + x) = x) 171 | proof_termH_termG x = let (Just ok) = compare_termH_termG x in ok 172 | -- RESULT : WORKS FOR ALL X ! 173 | 174 | 175 | -- ---------------------- 176 | -- TEST 5 THAT SHOULD WORK 177 | -- ---------------------- 178 | 179 | 180 | -- Test for (e1 + e2) + -e3 when e2 = e3. It then gives e1. 181 | -- Here : e1 - y 182 | -- e2 = ((3 + (0-2))+x) 183 | -- e3 = 1+x 184 | termJ : (x:ZZ) -> (y:ZZ)-> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x, y] ((y + ((3 + (0-2))+x)) + (-(1+x))) 185 | termJ x y = PlusG _ (PlusG _ (VarG _ _ (RealVariable _ _ _ _ (FS FZ))) 186 | (PlusG _ (PlusG _ (ConstG _ _ _ (Pos 3)) 187 | (MinusG _ (ConstG _ _ _ (Pos 0)) (ConstG _ _ _ (Pos 2)))) 188 | (VarG _ _ (RealVariable _ _ _ _ FZ)))) 189 | (NegG _ (PlusG _ (ConstG _ _ _ (Pos 1)) (VarG _ _ (RealVariable _ _ _ _ FZ)))) 190 | 191 | 192 | termK : (x:ZZ) -> (y:ZZ)-> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x, y] y 193 | termK x y = VarG _ _ (RealVariable _ _ _ _ (FS FZ)) 194 | 195 | termL : (x:ZZ) -> (y:ZZ)-> ExprG (%instance) (FakeSetAndMult (group_to_set (%instance))) [x, y] ((y + (1+x)) + (-(1+x))) 196 | termL x y = PlusG _ (PlusG _ (VarG _ _ (RealVariable _ _ _ _ (FS FZ))) 197 | (PlusG _ (ConstG _ _ _ (Pos 1)) 198 | (VarG _ _ (RealVariable _ _ _ _ FZ)))) 199 | (NegG _ (PlusG _ (ConstG _ _ _ (Pos 1)) (VarG _ _ (RealVariable _ _ _ _ FZ)))) 200 | 201 | compare_termJ_termK : (x:ZZ) -> (y:ZZ) -> Maybe (((y + ((3 + (0-2))+x)) + (-(1+x))) = y) 202 | compare_termJ_termK x y = groupDecideEq (%instance) (termJ x y) (termK x y) 203 | 204 | -- Later, we will have a real tactic "Group" which can fail... 205 | -- A quite complicated example 206 | proof_termJ_termK : (x:ZZ) -> (y:ZZ) -> (((y + ((3 + (0-2))+x)) + (-(1+x))) = y) 207 | proof_termJ_termK x y = let (Just ok) = compare_termJ_termK x y in ok 208 | -- RESULT : WORKS FOR ALL X AND Y ! 209 | 210 | 211 | -- Old stuff for debugging, no longer needed 212 | 213 | -- A) 214 | 215 | compare_termJ_termL : (x:ZZ) -> (y:ZZ) -> Maybe (((y + ((3 + (0-2))+x)) + (-(1+x))) = ((y + (1+x)) + (-(1+x)))) 216 | compare_termJ_termL x y = groupDecideEq (%instance) (termJ x y) (termL x y) 217 | 218 | -- Later, we will have a real tactic "Group" which can fail... 219 | proof_termJ_termL : (x:ZZ) -> (y:ZZ) -> (((y + ((3 + (0-2))+x)) + (-(1+x))) = ((y + (1+x)) + (-(1+x)))) 220 | proof_termJ_termL x y = let (Just ok) = compare_termJ_termL x y in ok 221 | -- OK 222 | 223 | -- B) 224 | 225 | 226 | 227 | -- get the "real value" 228 | {- 229 | getC : {c:Type} -> {p:dataTypes.Group c} -> {n:Nat} -> {g:Vect n c} -> (ExprG p g c1) -> c 230 | getC (ConstG _ p c1) = c1 231 | getC (PlusG _ e1 e2) = Plus (getC e1) (getC e2) 232 | getC (MinusG _ e1 e2) = Minus (getC e1) (getC e2) 233 | getC (NegG _ e) = Neg (getC e) 234 | getC (VarG _ p i b) = (index i g) 235 | -} 236 | 237 | 238 | {- 239 | termJ_norm : (x:ZZ) -> (y:ZZ) -> (n2 ** (g2 ** (c1 ** ((((ExprG {n=2} (%instance) [x, y] c1)), (((y + ((3 + (0-2))+x)) + (-(1+x))) = c1)), SubSet [x,y] g2)))) 240 | termJ_norm x y = groupReduce _ _ [x, y] (termJ x y) 241 | -} 242 | 243 | 244 | 245 | 246 | ---------- Proofs ---------- 247 | Provers.group_test.MPlusZZ_preserves_equiv_1 = proof 248 | intros 249 | rewrite p1 250 | rewrite p2 251 | exact Refl 252 | 253 | Provers.group_test.MNeg_preserves_equiv_1 = proof 254 | intros 255 | rewrite p 256 | exact Refl 257 | 258 | 259 | -------------------------------------------------------------------------------- /Provers/magma_reduce.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File magma_reduce.idr 4 | -- Normalize an expression reflecting an element in a magma 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.magma_reduce 8 | 9 | import Decidable.Equality 10 | import Data.Vect 11 | import Provers.dataTypes 12 | 13 | 14 | %access public export 15 | 16 | %default total 17 | 18 | 19 | -- Normalization 20 | magmaReduce : {c:Type} -> {n:Nat} -> {p:Magma c} -> {neg:c->c} -> {setAndMult:SetWithMult c (magma_to_set_class p)} -> {g:Vect n c} -> {c1:c} -> (ExprMa p neg setAndMult g c1) -> (c2 ** (ExprMa p neg setAndMult g c2, c1~=c2)) 21 | magmaReduce (ConstMa p neg setAndMult g const) = (_ ** (ConstMa p neg setAndMult g const, set_eq_undec_refl const)) 22 | magmaReduce (PlusMa neg setAndMult {g=g} (ConstMa p _ _ _ const1) (ConstMa p _ _ _ const2)) = (_ ** (ConstMa p neg setAndMult g (Plus const1 const2), set_eq_undec_refl (Plus const1 const2))) 23 | magmaReduce (PlusMa neg setAndMult e1 e2) = 24 | let (r_ih1 ** (e_ih1, p_ih1)) = (magmaReduce e1) in 25 | let (r_ih2 ** (e_ih2, p_ih2)) = (magmaReduce e2) in 26 | ((Plus r_ih1 r_ih2) ** (PlusMa neg setAndMult e_ih1 e_ih2, ?MmagmaReduce1)) 27 | magmaReduce (VarMa p neg setAndMult {c1=c1} v) = (_ ** (VarMa p neg setAndMult v, set_eq_undec_refl c1)) 28 | 29 | 30 | buildProofMagma : {c:Type} -> {n:Nat} -> (p:Magma c) -> (setAndNeg:SetWithNeg c (magma_to_set_class p)) -> {setAndMult:SetWithMult c (magma_to_set_class p)} -> {g:Vect n c} -> {x:c} -> {y:c} -> {c1:c} -> {c2:c} -> (ExprMa p (neg setAndNeg) setAndMult g c1) -> (ExprMa p (neg setAndNeg) setAndMult g c2) -> (x ~= c1) -> (y ~= c2) -> (Maybe (x~=y)) 31 | buildProofMagma p setAndNeg e1 e2 lp rp with (exprMa_eq p setAndNeg _ _ e1 e2) 32 | buildProofMagma p setAndNeg e1 e2 lp rp | Just e1_equiv_e2 = ?MbuildProofMagma 33 | buildProofMagma p setAndNeg e1 e2 lp rp | Nothing = Nothing 34 | 35 | 36 | magmaDecideEq : {c:Type} -> {n:Nat} -> (p:Magma c) -> (setAndNeg:SetWithNeg c (magma_to_set_class p)) -> {setAndMult:SetWithMult c (magma_to_set_class p)} -> {g:Vect n c} -> {x:c} -> {y:c} -> (ExprMa p (neg setAndNeg) setAndMult g x) -> (ExprMa p (neg setAndNeg) setAndMult g y) -> Maybe (x~=y) 37 | -- e1 is the left side, e2 is the right side 38 | magmaDecideEq p setAndNeg e1 e2 = 39 | let (r_e1 ** (e_e1, p_e1)) = magmaReduce e1 in 40 | let (r_e2 ** (e_e2, p_e2)) = magmaReduce e2 in 41 | buildProofMagma p setAndNeg e_e1 e_e2 p_e1 p_e2 42 | 43 | 44 | ---------- Proofs ---------- 45 | Provers.magma_reduce.MmagmaReduce1 = proof 46 | intros 47 | mrefine Plus_preserves_equiv 48 | exact p_ih1 49 | exact p_ih2 50 | 51 | Provers.magma_reduce.MbuildProofMagma = proof 52 | intros 53 | refine Just 54 | mrefine eq_preserves_eq 55 | exact c1 56 | exact c2 57 | exact lp 58 | exact rp 59 | exact e1_equiv_e2 60 | 61 | 62 | 63 | 64 | 65 | -------------------------------------------------------------------------------- /Provers/magma_test.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File magma_test.idr 4 | -- test the normalization for magma 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.magma_test 8 | 9 | import Data.Vect 10 | import Data.Fin 11 | import Provers.globalDef 12 | import Provers.dataTypes 13 | import Provers.magma_reduce 14 | 15 | 16 | %access public export 17 | 18 | implementation Set Nat where 19 | -- The relation is just the (syntactical) equality 20 | (~=) x y = (x = y) 21 | 22 | set_eq x y with (decEq x y) 23 | set_eq x x | Yes Refl = Just Refl 24 | set_eq x y | _ = Nothing 25 | 26 | -- proof that the relation of equality is an equivalence relation 27 | set_eq_undec_refl x = Refl 28 | set_eq_undec_sym p = sym p 29 | set_eq_undec_trans p1 p2 = rewrite p1 in rewrite p2 in Refl 30 | 31 | 32 | implementation Magma Nat where 33 | Plus x y = plus x y 34 | 35 | -- proof that this plus preserves the (syntactical) equality 36 | Plus_preserves_equiv p1 p2 = ?MPlus_preserves_equiv_1 37 | 38 | 39 | 40 | test1 : (x:Nat) -> ExprMa (%instance) (\x =>x) (FakeSetAndMult (%instance)) [x] (Plus 2 (Plus 3 x)) 41 | test1 x = PlusMa _ _ (ConstMa _ _ _ _ 2) (PlusMa _ _ (ConstMa _ _ _ _ 3) (VarMa _ _ _ (RealVariable _ _ _ _ FZ))) 42 | 43 | test2 : (x:Nat) -> ExprMa (%instance) (\x => x) (FakeSetAndMult (%instance)) [x] (Plus 5 x) 44 | test2 x = PlusMa _ _ (PlusMa _ _ (ConstMa _ _ _ _ 2) (ConstMa _ _ _ _ 3)) (VarMa _ _ _ (RealVariable _ _ _ _ FZ)) 45 | 46 | test3 : (x:Nat) -> ExprMa (%instance) (\x => x) (FakeSetAndMult (%instance)) [x] (Plus 5 x) 47 | test3 x = PlusMa _ _ (ConstMa _ _ _ _ 5) (VarMa _ _ _ (RealVariable _ _ _ _ FZ)) 48 | 49 | --First test : 2 + (3 + x) =\= 5 + x 50 | total -- cool ! 51 | compare_test1_test3 : (x:Nat) -> Maybe (2 + (3 + x) = 5 + x) 52 | compare_test1_test3 x = magmaDecideEq (%instance) (FakeSetAndNeg _) (test1 x) (test3 x) 53 | 54 | 55 | test1_not_equal_to_test3 : (x:Nat) -> (2 + (3 + x) = 5 + x) 56 | test1_not_equal_to_test3 x = let (Just pr) = magmaDecideEq (%instance) (FakeSetAndNeg _) (test1 x) (test3 x) in pr --A "non regression test", unfortunately not using the type checker (need to compute this term and to see if it crashs or not) 57 | -- Should crash if we use the value ! 58 | -- AND EFECTIVELY CRASHES FOR ALL X 59 | 60 | -- Second test : (2 + 3) + x = 5 + x 61 | total -- cool ! 62 | compare_test2_test3 : (x:Nat) -> Maybe ((2 + 3) + x = 5 + x) 63 | compare_test2_test3 x = magmaDecideEq (%instance) (FakeSetAndNeg _) (test2 x) (test3 x) 64 | 65 | test2_equal_test3 : (x:Nat) -> ((2 + 3) + x = 5 + x) 66 | test2_equal_test3 = \x => let (Just pr) = magmaDecideEq (%instance) (FakeSetAndNeg _) (test2 x) (test3 x) in pr --A second "non regression test", unfortunately not using the type checker (need to compute this term and to see if it crashs or not) 67 | -- WORKS FOR ALL X !! 68 | 69 | -- JUST A STUPID TEST TO UNDERSTAND WHAT HAPPEN IF I A CONSTANT IS IN FACT A VARIABLE 70 | -- Of course, it won't give the proof we want for all x (because the algorithm waits for the value of x since we're supposed to have a _constant_), 71 | -- but it works for specific values of x, which is what we would expect 72 | 73 | 74 | termX : (x:Nat) -> ExprMa (%instance) (\x => x) (FakeSetAndMult (%instance)) [x] x 75 | termX x = ConstMa _ _ _ _ x 76 | 77 | total -- cool ! 78 | compare_termX_termX : (x:Nat) -> Maybe (x = x) 79 | compare_termX_termX x = magmaDecideEq (%instance) (FakeSetAndNeg _) (termX x) (termX x) 80 | 81 | result_termX_termX : (x:Nat) -> (x = x) 82 | result_termX_termX x = let (Just pr) = magmaDecideEq (%instance) (FakeSetAndNeg _) (termX x) (termX x) in pr 83 | 84 | 85 | 86 | 87 | ---------- Proofs ---------- 88 | Provers.magma_test.MPlus_preserves_equiv_1 = proof 89 | intros 90 | rewrite p1 91 | rewrite p2 92 | exact Refl 93 | 94 | 95 | 96 | 97 | -------------------------------------------------------------------------------- /Provers/mathsResults.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File mathsResults.idr 4 | -- Various lemmas and theorems about algebraic structures, *not* strictly needed fot implementation of the ring tactic for Idris, 5 | -- but nice enough for being formulated here 6 | ---------------------------------------------------------------------------------------------------------------------------------- 7 | 8 | module Provers.mathsResults 9 | 10 | import Data.ZZ 11 | import Provers.globalDef 12 | import Provers.dataTypes 13 | import Provers.tools 14 | 15 | 16 | %default total 17 | 18 | -- --------------------------------- 19 | -- Mathematical results for a Group 20 | ------------------------------------ 21 | 22 | -- This is a (logical) DEFINITION, not a LEMMA (lies in TYPE) 23 | bad_push_negation : (C:Type) -> (p:dataTypes.Group C) -> (x:C) -> (y:C) -> Type 24 | bad_push_negation C p x y = (Neg (Plus x y) ~= Plus (Neg x) (Neg y)) 25 | 26 | bad_swap : (C:Type) -> (p:dataTypes.Group C) -> ((x:C) -> (y:C) -> (bad_push_negation C p x y)) -> ((x:C) -> (y:C) -> (Plus (Neg x) (Neg y) ~= Plus (Neg y) (Neg x))) 27 | bad_swap C p Hbad x y = 28 | let aux : (Neg (Plus x y) ~= Plus (Neg x) (Neg y)) = Hbad x y in 29 | let aux2 : (Neg (Plus x y) ~= Plus (Neg y) (Neg x)) = ?Mbad_swap_1 in 30 | ?Mbad_swap_2 31 | 32 | -- Proof that (forall x y, bad_push_negation x y) -> (forall x y, Plus x y ~= Plus y x) 33 | -- That is too say, if we could develop the negation in the left-right order, then any group would be a commutative group 34 | -- The conclusion is to *never* develop Neg (Plus x y) as Plus (Neg x) (Neg y) in our normalization procedure 35 | bad_commutativity : (C:Type) -> (p:dataTypes.Group C) -> ((x:C) -> (y:C) -> (bad_push_negation C p x y)) -> ((x:C) -> (y:C) -> Plus x y ~= Plus y x) 36 | bad_commutativity C p Hbad x y = 37 | let aux : (Neg (Neg x) ~= x) = group_doubleNeg C p x in 38 | let aux2 : (Neg (Neg y) ~= y) = group_doubleNeg C p y in 39 | let aux3 : (Plus x y ~= Plus (Neg (Neg x)) (Neg (Neg y))) = ?Mbad_commutativity_1 in 40 | let aux4 : (Plus (Neg (Neg x)) (Neg (Neg y)) ~= Plus (Neg (Neg y)) (Neg (Neg x))) = bad_swap C p Hbad (Neg x) (Neg y) in 41 | let aux5 : (Plus x y ~= Plus (Neg (Neg y)) (Neg (Neg x))) = ?Mbad_commutativity_2 in 42 | let aux6 : (Plus y x ~= Plus (Neg (Neg y)) (Neg (Neg x))) = ?Mbad_commutativity_3 in 43 | ?Mbad_commutativity_4 44 | 45 | {- 46 | -- Question : HOW CAN I DEFINE AN INSTANCE "ON THE FLY" ? 47 | bad_push_negation_IMPLIES_commutativeGroup : (C:Type) -> (p:dataTypes.Group C) -> ((x:C) -> (y:C) -> (bad_push_negation C p x y)) -> (dataTypes.CommutativeGroup C) 48 | bad_push_negation_IMPLIES_commutativeGroup C p Hbad = 49 | let comm : ((x:C) -> (y:C) -> (Plus x y ~= Plus y x)) = bad_commutativity C p Hbad in 50 | (instance CommutativeGroup C where 51 | Plus_comm c1 c2 = comm c1 c2) 52 | --?Mbad_push_negation_IMPLIES_commutativeGroup_1 53 | -} 54 | 55 | 56 | 57 | ---------- Proofs ---------- 58 | Provers.mathsResults.Mbad_swap_1 = proof 59 | intros 60 | exact (push_negation C p x y) 61 | 62 | Provers.mathsResults.Mbad_swap_2 = proof 63 | intros 64 | mrefine eq_preserves_eq 65 | exact (Plus (Neg x) (Neg y)) 66 | exact (Neg (Plus x y)) 67 | mrefine set_eq_undec_refl 68 | mrefine set_eq_undec_sym 69 | mrefine set_eq_undec_sym 70 | exact aux2 71 | exact aux 72 | 73 | Provers.mathsResults.Mbad_commutativity_1 = proof 74 | intros 75 | mrefine Plus_preserves_equiv 76 | mrefine set_eq_undec_sym 77 | mrefine set_eq_undec_sym 78 | exact aux 79 | exact aux2 80 | 81 | Provers.mathsResults.Mbad_commutativity_2 = proof 82 | intros 83 | mrefine eq_preserves_eq 84 | exact (Plus x y) 85 | exact (Plus (Neg (Neg x)) (Neg (Neg y))) 86 | mrefine set_eq_undec_refl 87 | mrefine set_eq_undec_sym 88 | exact aux3 89 | exact aux4 90 | 91 | Provers.mathsResults.Mbad_commutativity_3 = proof 92 | intros 93 | mrefine Plus_preserves_equiv 94 | mrefine set_eq_undec_sym 95 | mrefine set_eq_undec_sym 96 | exact aux2 97 | exact aux 98 | 99 | Provers.mathsResults.Mbad_commutativity_4 = proof 100 | intros 101 | mrefine eq_preserves_eq 102 | exact (Plus (Neg (Neg y)) (Neg (Neg x))) 103 | exact (Plus (Neg (Neg y)) (Neg (Neg x))) 104 | exact aux5 105 | exact aux6 106 | mrefine set_eq_undec_refl 107 | 108 | 109 | 110 | 111 | 112 | 113 | 114 | 115 | 116 | 117 | 118 | 119 | 120 | 121 | 122 | 123 | 124 | 125 | -------------------------------------------------------------------------------- /Provers/monoid_reduce.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File monoid_reduce.idr 4 | -- Normalize an expression reflecting an element in a monoid 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.monoid_reduce 8 | 9 | import Decidable.Equality 10 | import Data.Vect 11 | 12 | import Provers.dataTypes 13 | import Provers.semiGroup_reduce 14 | import Provers.tools 15 | 16 | %access public export 17 | 18 | --%default total 19 | 20 | 21 | 22 | total 23 | elimZero : (c:Type) -> (p:dataTypes.Monoid c) -> (neg:c->c) -> (setAndMult:SetWithMult c (monoid_to_set p)) -> (g:Vect n c) -> {c1:c} -> (ExprMo p neg setAndMult g c1) -> (c2 ** (ExprMo p neg setAndMult g c2, c1~=c2)) 24 | elimZero c p neg setAndMult g (ConstMo _ _ _ _ const) = (_ ** (ConstMo _ _ _ _ const, set_eq_undec_refl {c} _)) 25 | -- elimZero c p (PlusMo (ConstMo p const1) (VarMo p v)) = ?foo --with (monoid_get_setEq p Zero const1) 26 | elimZero c p neg setAndMult g (PlusMo _ _ (ConstMo _ _ _ _ const1) (VarMo _ _ _ v)) with (monoid_eq_as_elem_of_set p Zero const1) 27 | -- elimZero c p (PlusMo (ConstMo p const1) (VarMo p v)) | with_pat = ?elimZero_rhs~~ 28 | -- elimZero c p (PlusMo (ConstMo p const1) (VarMo p v)) | (Just prf) = ?elimZero_rhs_2 29 | elimZero c p neg setAndMult g (PlusMo _ _ (ConstMo _ _ _ _ const1) (VarMo _ _ _ v)) | (Just const1_eq_zero) = (_ ** (VarMo _ _ _ v, ?MelimZero1)) 30 | elimZero c p neg setAndMult g (PlusMo _ _ (ConstMo _ _ _ _ const1) (VarMo _ _ _ v)) | _ = (_ ** (PlusMo _ _ (ConstMo _ _ _ _ const1) (VarMo _ _ _ v), set_eq_undec_refl {c} _)) 31 | elimZero c p neg setAndMult g (PlusMo _ _ (VarMo _ _ _ v) (ConstMo _ _ _ _ const2)) with (monoid_eq_as_elem_of_set p Zero const2) 32 | elimZero c p neg setAndMult g (PlusMo _ _ (VarMo _ _ _ v) (ConstMo _ _ _ _ const2)) | (Just const2_eq_zero) = (_ ** (VarMo _ _ _ v, ?MelimZero2)) 33 | elimZero c p neg setAndMult g (PlusMo _ _ (VarMo _ _ _ v) (ConstMo _ _ _ _ const2)) | _ = (_ ** (PlusMo _ _ (VarMo _ _ _ v) (ConstMo _ _ _ _ const2), set_eq_undec_refl {c} _)) 34 | elimZero c p neg setAndMult g (PlusMo _ _ e1 e2) = 35 | let (r_ih1 ** (e_ih1, p_ih1)) = (elimZero c p neg setAndMult g e1) in 36 | let (r_ih2 ** (e_ih2, p_ih2)) = (elimZero c p neg setAndMult g e2) in 37 | ((Plus r_ih1 r_ih2) ** (PlusMo _ _ e_ih1 e_ih2, ?MelimZero3)) 38 | elimZero c p neg setAndMult g (VarMo _ _ _ v) = (_ ** (VarMo _ _ _ v, set_eq_undec_refl {c} _)) 39 | 40 | 41 | monoidReduce : (p:dataTypes.Monoid c) -> (setAndNeg:SetWithNeg c (monoid_to_set p)) -> {setAndMult:SetWithMult c (monoid_to_set p)} -> {g:Vect n c} -> {c1:c} -> (ExprMo p (neg setAndNeg) setAndMult g c1) -> (c2 ** (ExprMo p (neg setAndNeg) setAndMult g c2, c1~=c2)) 42 | monoidReduce p setAndNeg e = 43 | let (r_SGred ** (e_SGred, p_SGred)) = semiGroupReduce (monoid_to_semiGroup_class p) setAndNeg (monoid_to_semiGroup e) in 44 | let (r_elim ** (e_elim, p_elim)) = elimZero _ p _ _ _ (semiGroup_to_monoid p e_SGred) in 45 | (_ ** (e_elim, ?MmonoidReduce1)) 46 | 47 | 48 | total 49 | buildProofMonoid : (p:dataTypes.Monoid c) -> (setAndNeg:SetWithNeg c (monoid_to_set p)) -> {setAndMult:SetWithMult c (monoid_to_set p)} -> {g:Vect n c} -> {x : c} -> {y : c} -> (ExprMo p (neg setAndNeg) setAndMult g c1) -> (ExprMo p (neg setAndNeg) setAndMult g c2) -> (x ~= c1) -> (y ~= c2) -> (Maybe (x~=y)) 50 | buildProofMonoid p setAndNeg e1 e2 lp rp with (exprMo_eq p _ _ _ e1 e2) 51 | buildProofMonoid p setAndNeg e1 e2 lp rp | Just e1_equiv_e2 = ?MbuildProofMonoid 52 | buildProofMonoid p setAndNeg e1 e2 lp rp | Nothing = Nothing 53 | 54 | 55 | monoidDecideEq : (p:dataTypes.Monoid c) -> (setAndNeg:SetWithNeg c (monoid_to_set p)) -> {setAndMult:SetWithMult c (monoid_to_set p)} -> {g:Vect n c} -> (ExprMo p (neg setAndNeg) setAndMult g x) -> (ExprMo p (neg setAndNeg) setAndMult g y) -> (Maybe (x~=y)) 56 | -- e1 is the left side, e2 is the right side 57 | monoidDecideEq p setAndNeg e1 e2 = 58 | let (r_e1 ** (e_e1, p_e1)) = monoidReduce p setAndNeg e1 in 59 | let (r_e2 ** (e_e2, p_e2)) = monoidReduce p setAndNeg e2 in 60 | buildProofMonoid p setAndNeg e_e1 e_e2 p_e1 p_e2 61 | 62 | 63 | ---------- Proofs ---------- 64 | Provers.monoid_reduce.MmonoidReduce1 = proof 65 | intros 66 | mrefine eq_preserves_eq 67 | exact r_SGred 68 | exact r_SGred 69 | exact p_SGred 70 | mrefine set_eq_undec_sym 71 | mrefine set_eq_undec_refl 72 | exact p_elim 73 | 74 | Provers.monoid_reduce.MelimZero1 = proof 75 | intros 76 | mrefine eq_preserves_eq 77 | exact (Plus Zero c2) 78 | exact c2 79 | mrefine Plus_preserves_equiv 80 | mrefine set_eq_undec_refl 81 | mrefine Plus_neutral_1 82 | mrefine set_eq_undec_sym 83 | mrefine set_eq_undec_refl 84 | exact const1_eq_zero 85 | 86 | Provers.monoid_reduce.MelimZero2 = proof 87 | intros 88 | mrefine eq_preserves_eq 89 | exact (Plus c1 Zero) 90 | exact c1 91 | mrefine Plus_preserves_equiv 92 | mrefine set_eq_undec_refl 93 | mrefine Plus_neutral_2 94 | mrefine set_eq_undec_refl 95 | mrefine set_eq_undec_sym 96 | exact const2_eq_zero 97 | 98 | Provers.monoid_reduce.MelimZero3 = proof 99 | intros 100 | mrefine Plus_preserves_equiv 101 | exact p_ih1 102 | exact p_ih2 103 | 104 | Provers.monoid_reduce.MbuildProofMonoid = proof 105 | intros 106 | refine Just 107 | mrefine eq_preserves_eq 108 | exact c1 109 | exact c2 110 | exact lp 111 | exact rp 112 | exact e1_equiv_e2 113 | 114 | 115 | 116 | 117 | 118 | 119 | 120 | 121 | 122 | 123 | 124 | 125 | 126 | 127 | 128 | 129 | -------------------------------------------------------------------------------- /Provers/monoid_test.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File monoid_test.idr 4 | -- test the normalization for monoid 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.monoid_test 8 | 9 | import Data.Vect 10 | import Data.Fin 11 | 12 | import Provers.globalDef 13 | import Provers.dataTypes 14 | import Provers.tools 15 | import Provers.monoid_reduce 16 | import Provers.semiGroup_test 17 | import Provers.magma_test 18 | 19 | %access public export 20 | 21 | 22 | implementation dataTypes.Monoid Nat where 23 | Zero = Z 24 | 25 | Plus_neutral_1 c = ?M_Nat_Monoid_1 26 | 27 | 28 | Plus_neutral_2 Z = ?M_Nat_Monoid_2 29 | Plus_neutral_2 (S pc) = let px = Plus_neutral_2 pc in ?M_Nat_Monoid_3 30 | 31 | 32 | 33 | 34 | 35 | a : (x:Nat) -> ExprMo (%instance) (\x => x) (FakeSetAndMult (monoid_to_set (%instance))) [x] (2 + (x + 0)) 36 | a x = PlusMo _ _ (ConstMo _ _ _ _ 2) (PlusMo _ _ (VarMo _ _ _ (RealVariable _ _ _ _ FZ)) (ConstMo _ _ _ _ 0)) 37 | 38 | 39 | b : (x:Nat) -> ExprMo (%instance) (\x => x) (FakeSetAndMult (monoid_to_set (%instance))) [x] (2 + x) 40 | b x = PlusMo _ _ (ConstMo _ _ _ _ 2) (VarMo _ _ _ (RealVariable _ _ _ _ FZ)) 41 | 42 | 43 | -- Normalisation of 2 + (0 + x) that should give 2 + x, since now we are working on a monoid 44 | compare_a_b : (x:Nat) -> Maybe (2 + (x + 0) = 2 + x) 45 | compare_a_b x = monoidDecideEq (%instance) (FakeSetAndNeg _) (a x) (b x) 46 | 47 | -- Later, we will have a real tactic "Monoid" which can fail. At this point, we will 48 | -- not have a missing case for "Nothing", which enables now to manipulate some false proof 49 | -- (which causes a crash only when you apply then to a specific value for x) 50 | proof_a_b : (x:Nat) -> (2 + (x + 0) = 2 + x) 51 | proof_a_b x = let (Just ok) = compare_a_b x in ok 52 | -- WORKS FOR ALL X !! 53 | 54 | 55 | 56 | ---------- Proofs ---------- 57 | Provers.monoid_test.M_Nat_Monoid_1 = proof 58 | intro 59 | trivial 60 | 61 | Provers.monoid_test.M_Nat_Monoid_2 = proof 62 | trivial 63 | 64 | Provers.monoid_test.M_Nat_Monoid_3 = proof 65 | intros 66 | rewrite px 67 | trivial 68 | 69 | 70 | 71 | 72 | 73 | 74 | 75 | 76 | 77 | -------------------------------------------------------------------------------- /Provers/reflection.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File reflection.idr 4 | -- Reflect the concrete values into the abstract syntax 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.reflection 8 | 9 | 10 | import Decidable.Equality 11 | import Data.ZZ 12 | import Data.Fin 13 | import Data.Vect 14 | import Provers.dataTypes 15 | import Provers.tools 16 | import Provers.reflectionTools 17 | 18 | 19 | %access public export 20 | 21 | 22 | -- --------------------------------------------------------------------- 23 | -- GENERALISED VERSION GOING FROM ANY TYPE c (SEEN AS A RING) TO ExprR 24 | -- --------------------------------------------------------------------- 25 | 26 | %reflection 27 | reflectTermForRing : {c:Type} -> {n:Nat} -> (p:Ring c) -> (g : Vect n c) -> (reflectCst:TypeReflectConstants p) -> (x:c) -> (n' ** (g':Vect n' c ** (ExprR p (g ++ g') x))) 28 | reflectTermForRing {c=c} {n=n} p g reflectCst (Plus a b) with (reflectTermForRing p g reflectCst a) 29 | -- Reflecting sums... 30 | reflectTermForRing {c=c} {n=n} p g reflectCst (Plus a b) | (n' ** (g' ** a')) with (reflectTermForRing p (g ++ g') reflectCst b) 31 | reflectTermForRing {c=c} {n=n} p g reflectCst (Plus a b) | (n' ** (g' ** a')) | (n'' ** (g'' ** b')) = 32 | let this = PlusR {c=c} (weakenR g'' a') b' in 33 | ((n' + n'') ** ((g'++g'') ** (convertVectInExprR (plusAssociative n n' n'') (vectAppendAssociative g g' g'') this))) 34 | -- Reflecting substractions... 35 | reflectTermForRing {c=c} {n=n} p g reflectCst (Minus a b) with (reflectTermForRing p g reflectCst a) 36 | reflectTermForRing {c=c} {n=n} p g reflectCst (Minus a b) | (n' ** (g' ** a')) with (reflectTermForRing p (g ++ g') reflectCst b) 37 | reflectTermForRing {c=c} {n=n} p g reflectCst (Minus a b) | (n' ** (g' ** a')) | (n'' ** (g'' ** b')) = 38 | let this = MinusR {c=c} (weakenR g'' a') b' in 39 | ((n' + n'') ** ((g'++g'') ** (convertVectInExprR (plusAssociative n n' n'') (vectAppendAssociative g g' g'') this))) 40 | -- Reflecting opposites 41 | reflectTermForRing {c=c} {n=n} p g reflectCst (Neg a) with (reflectTermForRing p g reflectCst a) 42 | reflectTermForRing {c=c} {n=n} p g reflectCst (Neg a) | (n' ** (g' ** a')) = 43 | let this = NegR {c=c} a' in 44 | (n' ** (g' ** this)) 45 | -- Reflecting products.... 46 | reflectTermForRing {c=c} {n=n} p g reflectCst (Mult a b) with (reflectTermForRing p g reflectCst a) 47 | reflectTermForRing {c=c} {n=n} p g reflectCst (Mult a b) | (n' ** (g' ** a')) with (reflectTermForRing p (g ++ g') reflectCst b) 48 | reflectTermForRing {c=c} {n=n} p g reflectCst (Mult a b) | (n' ** (g' ** a')) | (n'' ** (g'' ** b')) = 49 | let this = MultR {c=c} (weakenR g'' a') b' in 50 | ((n' + n'') ** ((g'++g'') ** (convertVectInExprR (plusAssociative n n' n'') (vectAppendAssociative g g' g'') this))) 51 | -- Constants and variables 52 | reflectTermForRing {n=n} p g reflectCst t = 53 | case reflectCst g t of 54 | -- If reflectConstants has decided that this thing is a constant, then we trust it... 55 | Just this => ?MreflectTermForRing_1 -- I just return (Z ** ([] ** this)) with a few conversions for having the type aggreing (because n+0=n and g++[] = [] at the same time) 56 | -- If not, then it should be considered as a variable 57 | Nothing => case (isElement t g) of 58 | Just (i ** pr) => let this2 = VarR {n=n+Z} {g=g++Data.Vect.Nil} p (RealVariable {n=n+Z} _ _ _ (g++Data.Vect.Nil) (replace (sym (a_plus_zero n)) i)) in 59 | ?MreflectTermForRing_2 -- (Z ** (Data.VectType.Vect.Nil ** this)) 60 | Nothing => let this3 = VarR {n=n+S Z} {g=g++[t]} p (RealVariable {n=n+S Z} _ _ _ (g++[t]) (lastElement' n)) in 61 | ?MreflectTermForRing_3 -- (S Z ** ((t::Data.VectType.Vect.Nil) ** this)) 62 | 63 | 64 | 65 | -- CURRENT PROBLEMS WITH THE AUTOMATIC REFLECTION : 66 | -- ------------------------------------------------ 67 | -- There are two problems at the moment with the general automatic reflection for Ring (which is just above) : 68 | -- 1) It doesn't reflect the effective implementation of the abstract operations (Plus,Minus,Mult...) of the typeclass. 69 | -- for example, if you test (\xx,yy,uu,gg:ZZ => leftDep ( rightDep( reflectTermZForRing [] reflectConstantsForZ (xx+yy+uu+gg)))) then it goes 70 | -- directly in the last case, and it adds the whole formulae in the context. 71 | -- -> Idris needs to be improved for (really) allowing reflection with abstract operation that might become concrete with an actual instance. At the moment, 72 | -- idris allows it, but it doesn't work as expected 73 | 74 | -- 2) Problem with Idris's %reflection when calling subfunctions. See the issue on a smaller problem on Github here : https://github.com/idris-lang/Idris-dev/issues/2843 75 | 76 | 77 | 78 | ---------- Proofs ---------- 79 | 80 | Provers.reflection.MreflectTermForRing_1 = proof 81 | intros 82 | mrefine MkDPair 83 | exact Z 84 | compute 85 | mrefine MkDPair 86 | exact Nil 87 | compute 88 | mrefine convertVectInExprR 89 | exact n 90 | exact g 91 | mrefine plusZeroRightNeutral 92 | mrefine vectNilRightNeutral 93 | exact this 94 | 95 | Provers.reflection.MreflectTermForRing_2 = proof 96 | intros 97 | mrefine MkDPair 98 | exact Z 99 | compute 100 | mrefine MkDPair 101 | exact Nil 102 | compute 103 | rewrite pr 104 | rewrite (index_n_plus_0 g i) 105 | exact this2 106 | 107 | Provers.reflection.MreflectTermForRing_3 = proof 108 | intros 109 | mrefine MkDPair 110 | exact (S Z) 111 | compute 112 | mrefine MkDPair 113 | exact ([t]) 114 | compute 115 | exact (replace (indexOfLastElem g t) this3) 116 | 117 | 118 | 119 | 120 | 121 | -------------------------------------------------------------------------------- /Provers/reflectionForSpecificDataTypes.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File reflectionForSpecificDataTypes.idr 4 | -- Reflect the concrete values into the abstract syntax for specific datatypes (for testing purpose only) 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.reflectionForSpecificDataTypes 8 | 9 | import Decidable.Equality 10 | import Data.ZZ 11 | import Data.Fin 12 | import Data.Vect 13 | import Provers.dataTypes 14 | import Provers.tools 15 | import Provers.reflectionTools 16 | 17 | import Provers.ring_reduce -- Needed for the full_auto attempt at the end of the file 18 | 19 | import Provers.ring_test 20 | import Provers.commutativeGroup_test 21 | import Provers.group_test 22 | import Provers.monoid_test 23 | import Provers.semiGroup_test 24 | import Provers.magma_test 25 | 26 | %access public export 27 | 28 | 29 | -- ----------------------------------------------- 30 | -- REFLECTION FROM Nat (SEEN AS A MAGMA) TO ExprMa 31 | -- ----------------------------------------------- 32 | 33 | -- Reflects only from Nat (seen as a Magma) to ExprMa 34 | -- %logging 1 35 | %reflection 36 | reflectTermNat : {n:Nat} -> (g : Vect n Nat) -> (x:Nat) -> (n' ** (g':Vect n' Nat ** (ExprMa {c=Nat} {n=n+n'} (%instance) (neg (FakeSetAndNeg (magma_to_set_class (%instance)))) (FakeSetAndMult (magma_to_set_class (%instance))) (g ++ g') x))) 37 | reflectTermNat {n=n} g (a+b) with (reflectTermNat g a) 38 | reflectTermNat {n=n} g (a+b) | (n' ** (g' ** a')) with (reflectTermNat (g ++ g') b) 39 | reflectTermNat {n=n} g (a+b) | (n' ** (g' ** a')) | (n'' ** (g'' ** b')) = 40 | let this = PlusMa (neg (FakeSetAndNeg (magma_to_set_class (%instance)))) (FakeSetAndMult (magma_to_set_class (%instance))) (weakenMa g'' a') b' in 41 | ((n' + n'') ** ((g'++g'') ** (convertVectInExprMa (plusAssociative n n' n'') (vectAppendAssociative g g' g'') this))) 42 | -- Fix Idris : Huge problem if convertVectInExprMa was taking proofs that (n=n') and (g=g') (instead of the other way arround) and if I give the "sym" for these proofs here. 43 | -- %logging 0 44 | reflectTermNat {n=n} g t with (isElement t g) 45 | | Just (i ** p) = let this = VarMa {c=Nat} {n=n+Z} {g=g++Data.Vect.Nil} (%instance) (neg (FakeSetAndNeg (magma_to_set_class (%instance)))) (FakeSetAndMult (magma_to_set_class (%instance))) (RealVariable {n=n+Z} _ _ _ (g++Data.Vect.Nil) (replace (sym (a_plus_zero n)) i)) in 46 | ?MreflectTermNat_1 -- (Z ** (Data.VectType.Vect.Nil ** this)) 47 | | Nothing = let this = VarMa {c=Nat} {n=n+S Z} {g=g++[t]} (%instance) (neg (FakeSetAndNeg (magma_to_set_class (%instance)))) (FakeSetAndMult (magma_to_set_class (%instance))) (RealVariable {n=n+S Z} _ _ _ (g++[t]) (lastElement' n)) in 48 | ?MreflectTermNat_2 -- (S Z ** ((t::Data.VectType.Vect.Nil) ** this)) 49 | 50 | 51 | -- ---------------------------------------------- 52 | -- REFLECTION FROM ZZ (SEEN AS A RING) TO ExprR 53 | -- ---------------------------------------------- 54 | 55 | -- Reflects only from ZZ (seen as a Ring) to ExprR 56 | %reflection 57 | reflectTermZForRing : {n:Nat} -> (g : Vect n ZZ) -> (reflectCst:TypeReflectConstants {c=ZZ} (%instance)) -> (x:ZZ) -> (n' ** (g':Vect n' ZZ ** (ExprR {c=ZZ} {n=n+n'} (%instance) (g ++ g') x))) 58 | -- Reflecting sums... 59 | reflectTermZForRing {n=n} g reflectCst (a+b) with (reflectTermZForRing g reflectCst a) 60 | reflectTermZForRing {n=n} g reflectCst (a+b) | (n' ** (g' ** a')) with (reflectTermZForRing (g ++ g') reflectCst b) 61 | reflectTermZForRing {n=n} g reflectCst (a+b) | (n' ** (g' ** a')) | (n'' ** (g'' ** b')) = 62 | let this = PlusR (weakenR g'' a') b' in 63 | ((n' + n'') ** ((g'++g'') ** (convertVectInExprR (plusAssociative n n' n'') (vectAppendAssociative g g' g'') this))) 64 | -- Reflecting substractions... 65 | reflectTermZForRing {n=n} g reflectCst (a-b) with (reflectTermZForRing g reflectCst a) 66 | reflectTermZForRing {n=n} g reflectCst (a-b) | (n' ** (g' ** a')) with (reflectTermZForRing (g ++ g') reflectCst b) 67 | reflectTermZForRing {n=n} g reflectCst (a-b) | (n' ** (g' ** a')) | (n'' ** (g'' ** b')) = 68 | let this = MinusR (weakenR g'' a') b' in 69 | ((n' + n'') ** ((g'++g'') ** (convertVectInExprR (plusAssociative n n' n'') (vectAppendAssociative g g' g'') this))) 70 | -- Reflecting opposites 71 | reflectTermZForRing {n=n} g reflectCst (negate a) with (reflectTermZForRing g reflectCst a) 72 | reflectTermZForRing {n=n} g reflectCst (negate a) | (n' ** (g' ** a')) = 73 | let this = NegR a' in 74 | (n' ** (g' ** this)) 75 | -- Reflecting products.... 76 | reflectTermZForRing {n=n} g reflectCst (a*b) with (reflectTermZForRing g reflectCst a) 77 | reflectTermZForRing {n=n} g reflectCst (a*b) | (n' ** (g' ** a')) with (reflectTermZForRing (g ++ g') reflectCst b) 78 | reflectTermZForRing {n=n} g reflectCst (a*b) | (n' ** (g' ** a')) | (n'' ** (g'' ** b')) = 79 | let this = MultR (weakenR g'' a') b' in 80 | ((n' + n'') ** ((g'++g'') ** (convertVectInExprR (plusAssociative n n' n'') (vectAppendAssociative g g' g'') this))) 81 | -- Constants and variables 82 | reflectTermZForRing {n=n} g reflectCst t = 83 | case reflectCst g t of 84 | -- If reflectConstants has decided that this thing is a constant, then we trust it... 85 | Just this => ?MreflectTermZForRing_1 -- I just return (Z ** ([] ** this)) with a few conversions for having the type aggreing (because n+0=n and g++[] = [] at the same time) 86 | -- If not, then it should be considered as a variable 87 | Nothing => case (isElement t g) of 88 | Just (i ** pr) => let this2 = VarR {n=n+Z} {g=g++Data.Vect.Nil} (%instance) (RealVariable {n=n+Z} _ _ _ (g++Data.Vect.Nil) (replace (sym (a_plus_zero n)) i)) in 89 | ?MreflectTermZForRing_2 -- (Z ** (Data.Vect.Nil ** this)) 90 | Nothing => let this3 = VarR {n=n+S Z} {g=g++[t]} (%instance) (RealVariable {n=n+S Z} _ _ _ (g++[t]) (lastElement' n)) in 91 | ?MreflectTermZForRing_3 -- (S Z ** ((t::Data.VectType.Vect.Nil) ** this)) 92 | 93 | 94 | -- ------------------------------------------------------------------------------------------------- 95 | -- TEST OF THE REFLECTION MECHANISM FROM ZZ (SEEN AS A RING) TO ExprR : 96 | 97 | -- Let's try the automatic reflection 98 | -- With the example ((((((3*x)*(y*2))*u) + (x * (y-y))) + (3*((x*y)*(5*g)))) = (((3*x)*(y*5))*g) + (3*((x*y)*(2*u)))) 99 | -- that has been studied in ring_test.idr 100 | 101 | -- ------------------------------------------------------------------------------------------------- 102 | total 103 | %reflection 104 | reflectConstantsForZ : TypeReflectConstants {c=ZZ} (%instance) 105 | -- Fix Idris : I can't talk about {n=n} here (I don't need it though) 106 | reflectConstantsForZ g (Pos number) = Just (ConstR (%instance) g (Pos number)) 107 | reflectConstantsForZ g (NegS number) = Just (ConstR (%instance) g (NegS number)) 108 | reflectConstantsForZ g _ = Nothing 109 | 110 | 111 | -- Reflects only the LHS 112 | exprCr_automatic : (x:ZZ) -> (y:ZZ) -> (u:ZZ) -> (g:ZZ) -> (newG_size ** (newG:Vect newG_size ZZ ** (ExprR (%instance) newG (((((3*x)*(y*2))*u) + (x * (y-y))) + (3*((x*y)*(5*g))))))) 113 | exprCr_automatic x y u g = let (nAdded ** (gAdded ** exprReflected)) = reflectTermZForRing [] reflectConstantsForZ (((((3*x)*(y*2))*u) + (x * (y-y))) + (3*((x*y)*(5*g)))) in 114 | (_ ** (_ ** exprReflected)) 115 | 116 | 117 | -- ------------------------------------------------------------------------------------ 118 | -- For testing the reflection of the LHS : 119 | -- ---------------------------------------- 120 | -- 1) Showing the context that has been built automatically for the LHS : 121 | -- (\x,y,u,g => leftDep (rightDep (exprCr_automatic x y u g))) 122 | -- 2) Showing the reflected expression that has been built automatically for the LHS : 123 | -- (\x,y,u,g => rightDep (rightDep (exprCr_automatic x y u g))) 124 | -- ------------------------------------------------------------------------------------ 125 | 126 | -- Reflects only the RHS 127 | exprC2r_automatic : (x:ZZ) -> (y:ZZ) -> (u:ZZ) -> (g:ZZ) -> (newG_size ** (newG:Vect newG_size ZZ ** (ExprR (%instance) newG ((((3*x)*(y*5))*g) + (3*((x*y)*(2*u))))))) 128 | exprC2r_automatic x y u g = let (nAdded ** (gAdded ** exprReflected)) = reflectTermZForRing [] reflectConstantsForZ ((((3*x)*(y*5))*g) + (3*((x*y)*(2*u)))) in 129 | (_ ** (_ ** exprReflected)) 130 | 131 | -- ------------------------------------------------------------------------------------ 132 | -- For testing the reflection of the RHS : 133 | -- ---------------------------------------- 134 | -- 1) Showing the context that has been built automatically for the RHS : 135 | -- (\x,y,u,g => leftDep (rightDep (exprC2r_automatic x y u g))) 136 | -- 2) Showing the reflected expression thaa has been built automatically for the RHS : 137 | -- (\x,y,u,g => rightDep (rightDep (exprC2r_automatic x y u g))) 138 | -- ------------------------------------------------------------------------------------ 139 | 140 | -- Note : the context generated automatically for the LHS is [x,y,u,g], but the one for the RHS is [x,y,g,u] because of the order of appearance of the variables 141 | -- Update : that does not work anymore for the moment, because of the handler for constants only. I think there is a problem with Idris' %reflection machanism 142 | 143 | -- Now, let's try to do the reflection AND the proving, all in automatic mode... 144 | 145 | full_auto : (x:ZZ) -> (y:ZZ) -> (u:ZZ) -> (g:ZZ) -> ((((((3*x)*(y*2))*u) + (x * (y-y))) + (3*((x*y)*(5*g)))) = (((3*x)*(y*5))*g) + (3*((x*y)*(2*u)))) 146 | full_auto x y u g = let (nbAddedLHS ** (gammaAddedLHS ** lhs)) = reflectTermZForRing [] reflectConstantsForZ (((((3*x)*(y*2))*u) + (x * (y-y))) + (3*((x*y)*(5*g)))) in 147 | let (nbAddedRHS ** (gammaAddedRHS ** rhs)) = reflectTermZForRing gammaAddedLHS reflectConstantsForZ ((((3*x)*(y*5))*g) + (3*((x*y)*(2*u)))) in 148 | let (Just ok) = ringDecideEq (%instance) (weakenR gammaAddedRHS lhs) rhs in ok 149 | 150 | 151 | 152 | 153 | ---------- Proofs ---------- 154 | 155 | Provers.reflectionForSpecificDataTypes.MreflectTermNat_1 = proof 156 | intros 157 | mrefine MkDPair 158 | exact Z 159 | compute 160 | mrefine MkDPair 161 | exact Nil 162 | compute 163 | rewrite p 164 | rewrite (index_n_plus_0 g i) 165 | exact this 166 | 167 | Provers.reflectionForSpecificDataTypes.MreflectTermNat_2 = proof 168 | intros 169 | mrefine MkDPair 170 | exact (S Z) 171 | compute 172 | mrefine MkDPair 173 | exact ([t]) 174 | compute 175 | exact (replace (indexOfLastElem g t) this) 176 | 177 | Provers.reflectionForSpecificDataTypes.MreflectTermZForRing_1 = proof 178 | intros 179 | mrefine MkDPair 180 | exact Z 181 | compute 182 | mrefine MkDPair 183 | exact Nil 184 | compute 185 | mrefine convertVectInExprR 186 | exact n 187 | exact g 188 | mrefine plusZeroRightNeutral 189 | mrefine vectNilRightNeutral 190 | exact this 191 | 192 | Provers.reflectionForSpecificDataTypes.MreflectTermZForRing_2 = proof 193 | intros 194 | mrefine MkDPair 195 | exact Z 196 | compute 197 | mrefine MkDPair 198 | exact Nil 199 | compute 200 | rewrite pr 201 | rewrite (index_n_plus_0 g i) 202 | exact this2 203 | 204 | Provers.reflectionForSpecificDataTypes.MreflectTermZForRing_3 = proof 205 | intros 206 | mrefine MkDPair 207 | exact (S Z) 208 | compute 209 | mrefine MkDPair 210 | exact ([t]) 211 | compute 212 | exact (replace (indexOfLastElem g t) this3) 213 | 214 | 215 | 216 | 217 | -------------------------------------------------------------------------------- /Provers/reflectionTools.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File reflectionTools.idr 4 | -- Tools for the reflection (used in reflection.idr and in reflectionForSpecificDataTypes.idr) 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.reflectionTools 8 | 9 | import Decidable.Equality 10 | import Data.ZZ 11 | import Data.Fin 12 | import Data.Vect 13 | import Provers.dataTypes 14 | import Provers.tools 15 | 16 | 17 | %access public export 18 | 19 | 20 | 21 | -- -------------------------------------------------------------------------------------------------------------------------------------------------------------------------- 22 | -- Notes about the reflection mechanism : 23 | -- -------------------------------------- 24 | -- Edwin has switch from Fin to Elem in order to represent variables, which was a good idea, and which is particularly useful for the reflection machinery. 25 | -- Unfortunately, I've developped the entire collection of provers with Fin. Instead of changing completely the provers (that would be a nightmare as Fin is present everywhere !), 26 | -- I'll just adapt the automatic reflection, which becomes a bit more complicated. 27 | -- The key element of my reflection machinery is the usage of the function isElement, which returns an index AND the proof that this index effectively works. That enables 28 | -- to simulate the behaviour of Elem that I haven't used. I would say that my approach for the reflection is therefore the usual Coq approach, in the sense that I use more external lemmas. 29 | -- --------------------------------------------------------------------------------------------------------------------------------------------------------------------------- 30 | 31 | -- -------------------------------------- 32 | -- GENERAL REFLECTION TOOLS FOR RING --- 33 | -- -------------------------------------- 34 | total 35 | lemmaExtension : {c:Type} -> {n:Nat} -> {m:Nat} -> (g:Vect n c) -> (g':Vect m c) -> (i:Fin n) -> (index i g = index (convertFin _ i m) (g++g')) 36 | lemmaExtension Nil g' (FZ {k=k}) impossible 37 | lemmaExtension (gh::gt) g' (FZ {k=k}) = Refl 38 | lemmaExtension (gh::gt) g' (FS {k=Z} pi) = let proofOfFalse : Void = elimFinZero pi in -- Just elim the element of (Fin 0) that we have in the context to build an inhabitant of False (Void) 39 | ?MlemmaExtension_1 -- Just elim the inhabitant of False that we have constructed in the context 40 | lemmaExtension (gh::gt) g' (FS {k=S pk} pi) = let ih = lemmaExtension gt g' pi in ?MlemmaExtension_2 41 | 42 | 43 | 44 | total 45 | isElement : {a:Type} -> {n:Nat} -> (x : a) -> (G : Vect n a) -> Maybe (i:Fin n ** (index i G = x)) 46 | isElement x [] = Nothing 47 | isElement x (y :: ys) with (prim__syntactic_eq _ _ x y) 48 | isElement x (x :: ys) | Just Refl = Just (FZ ** Refl) -- [| Stop |] 49 | isElement x (y :: ys) | Nothing = let recCall = isElement x ys in 50 | -- [| Pop (isElem x ys) |] 51 | case recCall of 52 | Nothing => Nothing 53 | Just (i' ** p') => Just ((FS i') ** ?MisElement_1) 54 | 55 | 56 | 57 | -- --------------------------------------- 58 | -- SPECIFIC REFLECTION TOOLS FOR MAGMA --- 59 | -- --------------------------------------- 60 | 61 | -- NOT total : We don't treat if the variable is not a real variable, but we don't care since fake variables are only for encodings, and the user will never use them directly 62 | weakenMa : {c:Type} -> {p:Magma c} -> {n:Nat} -> {m:Nat} -> {g:Vect n c} -> (g':Vect m c) -> (ExprMa p (neg (FakeSetAndNeg (magma_to_set_class p))) (FakeSetAndMult (magma_to_set_class p)) g x) -> (ExprMa p (neg (FakeSetAndNeg (magma_to_set_class p))) (FakeSetAndMult (magma_to_set_class p)) (g ++ g') x) 63 | weakenMa g' (ConstMa _ _ _ g const1) = ConstMa _ _ _ (g++g') const1 64 | weakenMa g' (PlusMa _ _ e1 e2) = PlusMa _ _ (weakenMa g' e1) (weakenMa g' e2) 65 | weakenMa {n=n} {m=m} {g=g} g' (VarMa p _ _ (RealVariable _ _ _ _ i)) = 66 | let pExt = lemmaExtension g g' i in 67 | rewrite pExt in (VarMa {n=plus n m} p _ _ (RealVariable _ _ _ _ (convertFin _ i m))) 68 | 69 | 70 | total 71 | convertVectInExprMa : {c:Type} -> {p:Magma c} -> {n:Nat} -> {g:Vect n c} -> {x:c} -> 72 | {n':Nat} -> {g':Vect n' c} -> (n'=n) -> (g'=g) -> 73 | (ExprMa p (neg (FakeSetAndNeg (magma_to_set_class p))) (FakeSetAndMult (magma_to_set_class p)) g x) -> 74 | (ExprMa p (neg (FakeSetAndNeg (magma_to_set_class p))) (FakeSetAndMult (magma_to_set_class p)) g' x) 75 | convertVectInExprMa prEqN prEqG e with (prEqN) 76 | convertVectInExprMa prEqN prEqG e | (Refl) with (prEqG) 77 | convertVectInExprMa prEqN prEqG e | (Refl) | (Refl) = e -- Fix Idris : it works if I give directly e but if I put a metavariable, there's a bug when I try to prove it 78 | 79 | -- -------------------------------------- 80 | -- SPECIFIC REFLECTION TOOLS FOR RING --- 81 | -- -------------------------------------- 82 | 83 | -- NOT total : We don't treat if the variable is not a real variable, but we don't care since fake variables are only for encodings, and the user will never use them directly 84 | weakenR : {c:Type} -> {p:Ring c} -> {n:Nat} -> {m:Nat} -> {g:Vect n c} -> (g':Vect m c) -> (ExprR p g x) -> (ExprR p (g ++ g') x) 85 | weakenR g' (ConstR _ g const1) = ConstR _ (g++g') const1 86 | weakenR g' (PlusR e1 e2) = PlusR (weakenR g' e1) (weakenR g' e2) 87 | weakenR {n=n} {m=m} {g=g} g' (VarR p (RealVariable _ _ _ _ i)) = 88 | let pExt = lemmaExtension g g' i in 89 | rewrite pExt in (VarR {n=plus n m} p (RealVariable _ _ _ _ (convertFin _ i m))) 90 | 91 | 92 | total 93 | convertVectInExprR : {c:Type} -> {p:Ring c} -> {n:Nat} -> {g:Vect n c} -> {x:c} -> 94 | {n':Nat} -> {g':Vect n' c} -> (n'=n) -> (g'=g) -> 95 | (ExprR p g x) -> 96 | (ExprR p g' x) 97 | convertVectInExprR prEqN prEqG e with (prEqN) 98 | convertVectInExprR prEqN prEqG e | (Refl) with (prEqG) 99 | convertVectInExprR prEqN prEqG e | (Refl) | (Refl) = e -- Fix Idris : it works if I give directly e but if I put a metavariable, there's a bug when I try to prove it 100 | 101 | total 102 | TypeReflectConstants : {c:Type} -> (p:Ring c) -> Type 103 | TypeReflectConstants {c=c} p = ({n:Nat} -> (g:Vect n c) -> (x:c) -> Maybe (ExprR p g x)) 104 | 105 | 106 | ---------- Proofs ---------- 107 | 108 | Provers.reflectionTools.MlemmaExtension_1 = proof 109 | intros 110 | exact (void proofOfFalse) 111 | 112 | Provers.reflectionTools.MlemmaExtension_2 = proof 113 | intros 114 | rewrite (sym ih) 115 | rewrite (pre_convertFin_proofIrr pi _ (LTESucc (GTE_plus pk m)) (GTE_plus (S pk) m)) 116 | mrefine Refl 117 | 118 | Provers.reflectionTools.MisElement_1 = proof 119 | intros 120 | compute 121 | exact p' 122 | 123 | 124 | -------------------------------------------------------------------------------- /Provers/semiGroup_reduce.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File semiGroup_reduce.idr 4 | -- Normalize an expression reflecting an element in a semi-group 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.semiGroup_reduce 8 | 9 | import Decidable.Equality 10 | import Data.Vect 11 | 12 | import Provers.dataTypes 13 | import Provers.magma_reduce 14 | import Provers.tools 15 | 16 | %access public export 17 | 18 | 19 | --%default total 20 | 21 | 22 | -- Normalization 23 | -- No longer possible to tag this function as total due to fixed point to reach (non structural recursivity) (see last lines of the function) 24 | assoc : {c:Type} -> (p:SemiGroup c) -> (setAndNeg:SetWithNeg c (semiGroup_to_set p)) -> (setAndMult:SetWithMult c (semiGroup_to_set p)) -> (g:Vect n c) -> {c1:c} -> (ExprSG p (neg setAndNeg) setAndMult g c1) -> (c2 ** (ExprSG p (neg setAndNeg) setAndMult g c2, c1~=c2)) 25 | assoc p setAndNeg setAndMult g (ConstSG _ _ _ _ const) = (_ ** (ConstSG _ _ _ _ const, set_eq_undec_refl const)) 26 | assoc {c} p setAndNeg setAndMult g (VarSG _ _ _ v) = (_ ** (VarSG _ _ _ v, set_eq_undec_refl {c=c} _)) 27 | -- (x + c1) + (c2 + y) -> (x + (res c1+c2)) + y 28 | assoc p setAndNeg setAndMult g (PlusSG _ _ (PlusSG _ _ e1 (ConstSG _ _ _ _ const1)) (PlusSG _ _ (ConstSG _ _ _ _ const2) e2)) = 29 | let (r_ih1 ** (e_ih1, p_ih1)) = (assoc p setAndNeg setAndMult g e1) in 30 | let (r_ih2 ** (e_ih2, p_ih2)) = (assoc p setAndNeg setAndMult g e2) in 31 | let (r_3 ** (e_3, p_3)) = magmaReduce (semiGroup_to_magma {p=p} {neg=neg setAndNeg} {setAndMult=setAndMult} {g=g} (PlusSG _ _ (ConstSG _ _ _ _ const1) (ConstSG _ _ _ _ const2))) in 32 | let e_3' = magma_to_semiGroup p e_3 in 33 | (_ ** ((PlusSG _ _ (PlusSG _ _ e_ih1 e_3') e_ih2), ?Massoc1)) 34 | -- (x + c1) + c2 -> x + (res c1+c2) 35 | assoc p setAndNeg setAndMult g (PlusSG _ _ (PlusSG _ _ e1 (ConstSG _ _ _ _ const1)) (ConstSG _ _ _ _ const2)) = 36 | let (r_ih1 ** (e_ih1, p_ih1)) = (assoc p setAndNeg setAndMult g e1) in 37 | let (r_2 ** (e_2, p_2)) = magmaReduce (semiGroup_to_magma {p=p} {neg=neg setAndNeg} {setAndMult=setAndMult} {g=g} (PlusSG _ _ (ConstSG _ _ _ _ const1) (ConstSG _ _ _ _ const2))) in 38 | let e_2' = magma_to_semiGroup p e_2 in 39 | (_ ** ((PlusSG _ _ e_ih1 e_2'), ?Massoc2)) 40 | -- c1 + (c2 + x) -> (res c1 + c2) + x 41 | assoc p setAndNeg setAndMult g (PlusSG _ _ (ConstSG _ _ _ _ const1) (PlusSG _ _ (ConstSG _ _ _ _ const2) e1)) = 42 | let (r_ih1 ** (e_ih1, p_ih1)) = (assoc p setAndNeg setAndMult g e1) in 43 | let (r_2 ** (e_2, p_2)) = magmaReduce (semiGroup_to_magma {p=p} {neg=neg setAndNeg} {setAndMult} {g=g} (PlusSG _ _ (ConstSG _ _ _ _ const1) (ConstSG _ _ _ _ const2))) in 44 | let e_2' = magma_to_semiGroup p e_2 in 45 | (_ ** ((PlusSG _ _ e_2' e_ih1), ?Massoc3)) 46 | assoc p setAndNeg setAndMult g (PlusSG _ _ e1 e2) = 47 | let (r_ih1 ** (e_ih1, p_ih1)) = (assoc p setAndNeg setAndMult g e1) in 48 | let (r_ih2 ** (e_ih2, p_ih2)) = (assoc p setAndNeg setAndMult g e2) in 49 | let (r_3 ** (e_3, p_3)) = magmaReduce (semiGroup_to_magma {p=p} {neg=neg setAndNeg} {setAndMult} {g=g} e1) in 50 | let (r_4 ** (e_4, p_4)) = magmaReduce (semiGroup_to_magma {p=p} {neg=neg setAndNeg} {setAndMult} {g=g} e2) in 51 | let e_3' = magma_to_semiGroup p e_3 in 52 | let e_4' = magma_to_semiGroup p e_4 in 53 | case (exprSG_eq p setAndNeg setAndMult g (PlusSG _ _ e1 e2) (PlusSG _ _ e_3' e_4')) of 54 | Just _ => (_ ** ((PlusSG _ _ e_3' e_4'), ?Massoc4)) -- Fixed point reached 55 | Nothing => let (r_final ** (e_final, p_final)) = assoc p setAndNeg setAndMult g (PlusSG _ _ e_3' e_4') in -- Need to continue 56 | (_ ** (e_final, ?Massoc5)) 57 | 58 | 59 | total 60 | addAfter : {c:Type} -> (p:SemiGroup c) -> (setAndNeg:SetWithNeg c (semiGroup_to_set p)) -> (setAndMult:SetWithMult c (semiGroup_to_set p)) -> (g:Vect n c) -> {c1:c} -> {c2:c} -> (ExprSG p (neg setAndNeg) setAndMult g c1) -> (ExprSG p (neg setAndNeg) setAndMult g c2) -> (c3 ** (ExprSG p (neg setAndNeg) setAndMult g c3, c3~=Plus c1 c2)) 61 | addAfter {c} p setAndNeg setAndMult g (ConstSG _ _ _ _ const1) e = (_ ** (PlusSG _ _ (ConstSG _ _ _ _ const1) e, set_eq_undec_refl {c} _)) 62 | addAfter {c} p setAndNeg setAndMult g (VarSG _ _ _ v) e = (_ ** (PlusSG _ _ (VarSG _ _ _ v) e, set_eq_undec_refl {c} _)) 63 | addAfter p setAndNeg setAndMult g (PlusSG _ _ e11 e12) e2 = 64 | let (r_ih1 ** (e_ih1, p_ih1)) = addAfter p setAndNeg setAndMult g e12 e2 65 | in (_ ** (PlusSG _ _ e11 e_ih1, ?MaddAfter1)) 66 | 67 | 68 | -- Transforms an expression in the form x + (y + (z + ...)) 69 | -- can't be tagged as total (non structural recursion) 70 | shuffleRight : {c:Type} -> (p:SemiGroup c) -> (setAndNeg:SetWithNeg c (semiGroup_to_set p)) -> (setAndMult:SetWithMult c (semiGroup_to_set p)) -> (g:Vect n c) -> {c1:c} -> (ExprSG p (neg setAndNeg) setAndMult g c1) -> (c2 ** (ExprSG p (neg setAndNeg) setAndMult g c2, c1~=c2)) 71 | shuffleRight {c} p setAndNeg setAndMult g (ConstSG _ _ _ _ const1) = (_ ** (ConstSG _ _ _ _ const1, set_eq_undec_refl {c} _)) 72 | shuffleRight {c} p setAndNeg setAndMult g (VarSG _ _ _ v) = (_ ** (VarSG _ _ _ v, set_eq_undec_refl {c} _)) 73 | 74 | shuffleRight {c} p setAndNeg setAndMult g (PlusSG _ _ (ConstSG _ _ _ _ const1) (ConstSG _ _ _ _ const2)) = (_ ** (PlusSG _ _ (ConstSG _ _ _ _ const1) (ConstSG _ _ _ _ const2), set_eq_undec_refl {c} _)) 75 | shuffleRight {c} p setAndNeg setAndMult g (PlusSG _ _ (ConstSG _ _ _ _ const1) (VarSG _ _ _ v)) = (_ ** (PlusSG _ _ (ConstSG _ _ _ _ const1) (VarSG _ _ _ v), set_eq_undec_refl {c} _)) 76 | shuffleRight {c} p setAndNeg setAndMult g (PlusSG _ _ (ConstSG _ _ _ _ const1) (PlusSG _ _ e21 e22)) = 77 | let (r_ih1 ** (e_ih1, p_ih1)) = shuffleRight p setAndNeg setAndMult g (PlusSG _ _ e21 e22) in 78 | (_ ** (PlusSG _ _ (ConstSG _ _ _ _ const1) e_ih1, ?MshuffleRight1)) 79 | -- Previously : PlusSG (ConstSG c1) (addAfter (shuffleRight p21) (shuffleRight p22)) 80 | 81 | shuffleRight {c} p setAndNeg setAndMult g (PlusSG _ _ (VarSG _ _ _ v1) (ConstSG _ _ _ _ const2)) = (_ ** (PlusSG _ _ (VarSG _ _ _ v1) (ConstSG _ _ _ _ const2), set_eq_undec_refl {c} _)) 82 | shuffleRight {c} p setAndNeg setAndMult g (PlusSG _ _ (VarSG _ _ _ v1) (VarSG _ _ _ v2)) = (_ ** (PlusSG _ _ (VarSG _ _ _ v1) (VarSG _ _ _ v2), set_eq_undec_refl {c} _)) 83 | shuffleRight p setAndNeg setAndMult g (PlusSG _ _ (VarSG _ _ _ v1) (PlusSG _ _ e21 e22)) = 84 | let (r_ih1 ** (e_ih1, p_ih1)) = shuffleRight p setAndNeg setAndMult g (PlusSG _ _ e21 e22) in 85 | (_ ** (PlusSG _ _ (VarSG _ _ _ v1) e_ih1, ?MshuffleRight2)) 86 | -- PlusSG (VarSG v1) (addAfter (shuffleRight p21) (shuffleRight p22)) -- ok with me 87 | 88 | shuffleRight p setAndNeg setAndMult g (PlusSG _ _ (PlusSG _ _ e11 e12) (ConstSG _ _ _ _ const2)) = 89 | let (r_ih1 ** (e_ih1, p_ih1)) = shuffleRight p setAndNeg setAndMult g (PlusSG _ _ e11 e12) in 90 | let (r_2 ** (e_2, p_2)) = addAfter p setAndNeg setAndMult g e_ih1 (ConstSG _ _ _ _ const2) in 91 | (_ ** (e_2, ?MshuffleRight3)) 92 | shuffleRight p setAndNeg setAndMult g (PlusSG _ _ (PlusSG _ _ e11 e12) (VarSG _ _ _ v2)) = 93 | let (r_ih1 ** (e_ih1, p_ih1)) = shuffleRight p setAndNeg setAndMult g (PlusSG _ _ e11 e12) in 94 | let (r_2 ** (e_2, p_2)) = addAfter p setAndNeg setAndMult g e_ih1 (VarSG _ _ _ v2) in 95 | (_ ** (e_2, ?MshuffleRight4)) 96 | shuffleRight p setAndNeg setAndMult g (PlusSG _ _ (PlusSG _ _ e11 e12) (PlusSG _ _ e21 e22)) = 97 | let (r_ih1 ** (e_ih1, p_ih1)) = shuffleRight p setAndNeg setAndMult g (PlusSG _ _ e11 e12) in 98 | let (r_ih2 ** (e_ih2, p_ih2)) = shuffleRight p setAndNeg setAndMult g (PlusSG _ _ e21 e22) in 99 | let (r_3 ** (e_3, p_3)) = addAfter p setAndNeg setAndMult g e_ih1 e_ih2 in 100 | (_ ** (e_3, ?MshuffleRight5)) 101 | -- Previously : addAfter (addAfter (shuffleRight p11) (shuffleRight p12)) (addAfter (shuffleRight p21) (shuffleRight p22)) 102 | -- Note : equivalent to "addAfter (addAfter (addAfter (shuffleRight p11) (shuffleRight p12)) (shuffleRight p21)) (shuffleRight p22)" 103 | 104 | 105 | semiGroupReduce : (p:SemiGroup c) -> (setAndNeg:SetWithNeg c (semiGroup_to_set p)) -> {setAndMult:SetWithMult c (semiGroup_to_set p)} -> {g:Vect n c} -> {c1:c} -> (ExprSG p (neg setAndNeg) setAndMult g c1) -> (c2 ** (ExprSG p (neg setAndNeg) setAndMult g c2, c1~=c2)) 106 | semiGroupReduce p setAndNeg e = 107 | let (r_assoc ** (e_assoc, p_assoc)) = assoc p setAndNeg _ _ e in 108 | let (r_shuffle ** (e_shuffle, p_shuffle)) = shuffleRight p _ _ _ e_assoc in 109 | (_ ** (e_shuffle, ?MsemiGroupReduce1)) 110 | 111 | 112 | total 113 | buildProofSemiGroup : (p:SemiGroup c) -> (setAndNeg:SetWithNeg c (semiGroup_to_set p)) -> {setAndMult:SetWithMult c (semiGroup_to_set p)} -> {g:Vect n c} -> {x : c} -> {y : c} -> {c1:c} -> {c2:c} -> (ExprSG p (neg setAndNeg) setAndMult g c1) -> (ExprSG p (neg setAndNeg) setAndMult g c2) -> (x ~= c1) -> (y ~= c2) -> (Maybe (x~=y)) 114 | buildProofSemiGroup p setAndNeg e1 e2 lp rp with (exprSG_eq p _ _ _ e1 e2) 115 | buildProofSemiGroup p setAndNeg e1 e2 lp rp | Just e1_equiv_e2 = ?MbuildProofSemiGroup 116 | buildProofSemiGroup p setAndNeg e1 e2 lp rp | Nothing = Nothing 117 | 118 | 119 | semiGroupDecideEq : (p:SemiGroup c) -> (setAndNeg:SetWithNeg c (semiGroup_to_set p)) -> {setAndMult:SetWithMult c (semiGroup_to_set p)} -> {g:Vect n c} -> (ExprSG p (neg setAndNeg) setAndMult g x) -> (ExprSG p (neg setAndNeg) setAndMult g y) -> (Maybe (x~=y)) 120 | -- e1 is the left side, e2 is the right side 121 | semiGroupDecideEq p setAndNeg e1 e2 = 122 | let (r_e1 ** (e_e1, p_e1)) = semiGroupReduce p setAndNeg e1 in 123 | let (r_e2 ** (e_e2, p_e2)) = semiGroupReduce p setAndNeg e2 in 124 | buildProofSemiGroup p setAndNeg e_e1 e_e2 p_e1 p_e2 125 | 126 | 127 | 128 | ---------- Proofs ---------- 129 | -- NOTE : Idris is doing a strange job when proving the goal G by using something which requires you to prove the goal G' (ie, you've used G' -> G). 130 | -- Instead of immediately having to prove G' (ie, G' becomes at the top of the stack of things remaining to be proven), you will have to prove G' after all the other waiting subgoals 131 | Provers.semiGroup_reduce.Massoc1 = proof 132 | intros 133 | mrefine eq_preserves_eq 134 | exact (Plus (Plus c1 const1) (Plus const2 c2)) 135 | exact (Plus (Plus c1 (Plus const1 const2)) c2) 136 | mrefine set_eq_undec_refl 137 | mrefine Plus_preserves_equiv 138 | mrefine set_eq_undec_sym 139 | mrefine Plus_preserves_equiv 140 | mrefine set_eq_undec_sym 141 | exact (semiGroupAssoc_4terms c p c1 const1 const2 c2) 142 | mrefine set_eq_undec_sym 143 | mrefine set_eq_undec_sym 144 | exact p_ih2 145 | exact p_ih1 146 | exact p_3 147 | 148 | Provers.semiGroup_reduce.Massoc2 = proof 149 | intros 150 | mrefine eq_preserves_eq 151 | exact (Plus c1 (Plus const1 const2)) 152 | exact (Plus r_ih1 r_2) 153 | mrefine Plus_assoc 154 | mrefine set_eq_undec_refl 155 | mrefine Plus_preserves_equiv 156 | exact p_ih1 157 | exact p_2 158 | 159 | Provers.semiGroup_reduce.Massoc3 = proof 160 | intros 161 | mrefine eq_preserves_eq 162 | exact (Plus (Plus const1 const2) c2) 163 | exact (Plus r_2 r_ih1) 164 | mrefine set_eq_undec_sym 165 | mrefine set_eq_undec_refl 166 | mrefine Plus_preserves_equiv 167 | mrefine Plus_assoc 168 | exact p_2 169 | exact p_ih1 170 | 171 | Provers.semiGroup_reduce.Massoc4 = proof { 172 | intros 173 | trivial 174 | } 175 | 176 | Provers.semiGroup_reduce.Massoc5 = proof 177 | intros 178 | mrefine eq_preserves_eq 179 | exact (Plus r_3 r_4) 180 | exact r_final 181 | mrefine Plus_preserves_equiv 182 | mrefine set_eq_undec_sym 183 | exact p_final 184 | exact p_3 185 | exact p_4 186 | mrefine set_eq_undec_refl 187 | 188 | Provers.semiGroup_reduce.MaddAfter1 = proof 189 | intros 190 | mrefine eq_preserves_eq 191 | exact (Plus (Plus c1 c2) c3) 192 | exact (Plus (Plus c1 c2) c3) 193 | mrefine eq_preserves_eq 194 | mrefine set_eq_undec_refl 195 | mrefine set_eq_undec_refl 196 | exact (Plus c1 (Plus c2 c3)) 197 | exact (Plus (Plus c1 c2) c3) 198 | mrefine Plus_preserves_equiv 199 | mrefine set_eq_undec_refl 200 | mrefine set_eq_undec_sym 201 | mrefine set_eq_undec_refl 202 | exact p_ih1 203 | exact (Plus_assoc c1 c2 c3) 204 | 205 | Provers.semiGroup_reduce.MshuffleRight1 = proof 206 | intros 207 | mrefine Plus_preserves_equiv 208 | mrefine set_eq_undec_refl 209 | exact p_ih1 210 | 211 | Provers.semiGroup_reduce.MshuffleRight2 = proof 212 | intros 213 | mrefine Plus_preserves_equiv 214 | mrefine set_eq_undec_refl 215 | exact p_ih1 216 | 217 | Provers.semiGroup_reduce.MshuffleRight3 = proof 218 | intros 219 | mrefine eq_preserves_eq 220 | exact (Plus (Plus c1 c2) const2) 221 | exact (Plus (Plus c1 c2) const2) 222 | mrefine set_eq_undec_refl 223 | mrefine eq_preserves_eq 224 | mrefine set_eq_undec_refl 225 | exact (Plus r_ih1 const2) 226 | exact (Plus r_ih1 const2) 227 | exact p_2 228 | mrefine Plus_preserves_equiv 229 | mrefine set_eq_undec_refl 230 | mrefine set_eq_undec_sym 231 | mrefine set_eq_undec_refl 232 | mrefine set_eq_undec_sym 233 | exact p_ih1 234 | 235 | Provers.semiGroup_reduce.MshuffleRight4 = proof 236 | intros 237 | mrefine eq_preserves_eq 238 | exact (Plus r_ih1 c3) 239 | exact (Plus r_ih1 c3) 240 | mrefine Plus_preserves_equiv 241 | exact p_2 242 | mrefine set_eq_undec_refl 243 | exact p_ih1 244 | mrefine set_eq_undec_refl 245 | 246 | Provers.semiGroup_reduce.MshuffleRight5 = proof 247 | intros 248 | mrefine eq_preserves_eq 249 | exact (Plus r_ih1 r_ih2) 250 | exact (Plus r_ih1 r_ih2) 251 | mrefine Plus_preserves_equiv 252 | exact p_3 253 | mrefine set_eq_undec_refl 254 | exact p_ih1 255 | exact p_ih2 256 | 257 | Provers.semiGroup_reduce.MsemiGroupReduce1 = proof 258 | intros 259 | mrefine eq_preserves_eq 260 | exact r_assoc 261 | exact r_assoc 262 | exact p_assoc 263 | mrefine set_eq_undec_sym 264 | mrefine set_eq_undec_refl 265 | exact p_shuffle 266 | 267 | Provers.semiGroup_reduce.MbuildProofSemiGroup = proof 268 | intros 269 | refine Just 270 | mrefine eq_preserves_eq 271 | exact c1 272 | exact c2 273 | exact lp 274 | exact rp 275 | exact e1_equiv_e2 276 | 277 | 278 | 279 | 280 | 281 | 282 | 283 | 284 | 285 | 286 | 287 | 288 | 289 | 290 | 291 | 292 | 293 | 294 | 295 | 296 | 297 | -------------------------------------------------------------------------------- /Provers/semiGroup_test.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | -- File semiGroup_test.idr 4 | -- test the normalization for semiGroup 5 | ------------------------------------------------------------------- 6 | 7 | module Provers.semiGroup_test 8 | 9 | import Data.Vect 10 | import Data.Fin 11 | import Provers.globalDef 12 | import Provers.dataTypes 13 | import Provers.tools 14 | import Provers.semiGroup_reduce 15 | import Provers.magma_test 16 | 17 | %access public export 18 | 19 | implementation SemiGroup Nat where 20 | Plus_assoc c1 c2 c3 = sym (plusAssociative c1 c2 c3) 21 | 22 | 23 | test1' : (x:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x] (2 + (3 + x)) 24 | test1' x = PlusSG _ _ (ConstSG _ _ _ _ 2) (PlusSG _ _ (ConstSG _ _ _ _ 3) (VarSG _ _ _ (RealVariable _ _ _ _ FZ))) 25 | 26 | test2' : (x:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x] (5 + x) 27 | test2' x = PlusSG _ _ (PlusSG _ _ (ConstSG _ _ _ _ 2) (ConstSG _ _ _ _ 3)) (VarSG _ _ _ (RealVariable _ _ _ _ FZ)) 28 | 29 | test3' : (x:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x] (5 + x) 30 | test3' x = PlusSG _ _ (ConstSG _ _ _ _ 5) (VarSG _ _ _ (RealVariable _ _ _ _ FZ)) 31 | 32 | test4' : (x:Nat) -> (y:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x, y] ((x + (1+1)) + (2 + y)) 33 | test4' x y = (PlusSG _ _ (PlusSG _ _ (VarSG _ _ _ (RealVariable _ _ _ _ FZ)) 34 | (PlusSG _ _ (ConstSG _ _ _ _ 1) (ConstSG _ _ _ _ 1))) 35 | (PlusSG _ _ (ConstSG _ _ _ _ 2) (VarSG _ _ _ (RealVariable _ _ _ _ (FS FZ))))) 36 | 37 | test5' : (x:Nat) -> (y:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x, y] (x + (4 + y)) 38 | test5' x y = PlusSG _ _ (VarSG _ _ _ (RealVariable _ _ _ _ FZ)) (PlusSG _ _ (ConstSG _ _ _ _ 4) (VarSG _ _ _ (RealVariable _ _ _ _ (FS FZ)))) 39 | 40 | -- Normalisation of 2 + (3 + x) that should give 5 + x this time, since now we are working with semiGroup 41 | compare_test1'_test3' : (x:Nat) -> Maybe (2 + (3 + x) = 5 + x) 42 | compare_test1'_test3' x = semiGroupDecideEq (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (test1' x) (test3' x) 43 | 44 | -- Later, we will have a real tactic which can fail. At this point, we will 45 | -- not have a missing case for "Nothing", which enables now to manipulate some false proof 46 | -- (which causes a crash only when you apply then to a specific value for x) 47 | test : (x:Nat) -> (2 + (3 + x) = 5 + x) 48 | test x = let (Just ok) = compare_test1'_test3' x in ok 49 | -- WORKS FOR ALL X !! 50 | 51 | 52 | -- SECOND TEST : WE NORMALIZE TEST4' AND TEST5' 53 | 54 | {- 55 | get_r : {pr: SemiGroup c} -> {r1:c} -> (r ** (ExprSG pr neg [x, y] r, r1=r)) -> c 56 | get_r (r ** (e, p)) = r 57 | 58 | pre_get_e : {pr: SemiGroup c} -> {r1:c} -> (r ** (ExprSG pr neg [x, y] r, r1=r)) -> ExprSG pr neg [x, y] r 59 | pre_get_e (r ** (e, p)) = e 60 | 61 | get_e : {pr: SemiGroup c} -> {r1:c} -> (big:(r ** (ExprSG pr neg [x, y] r, r1=r))) -> ExprSG pr neg [x, y] (get_r big) 62 | get_e (r ** (e, p)) = e 63 | -} 64 | 65 | 66 | -- Result of the automatic equality solver for test4' and test5' 67 | secondTest : (x:Nat) -> (y:Nat) -> (((x + (1+1)) + (2 + y)) = (x + (4 + y))) 68 | secondTest x y = let (Just ok) = semiGroupDecideEq (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (test4' x y) (test5' x y) in ok 69 | -- RESULT : WORKS FOR ALL X AND Y ! 70 | 71 | 72 | -- Code to debug secondTest 73 | print_test4'_norm : Nat -> Nat -> String 74 | print_test4'_norm = (\x => \y => print_ExprSG show (left (rightDep (semiGroupReduce (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (test4' x y))))) 75 | 76 | print_test5'_norm : Nat -> Nat -> String 77 | print_test5'_norm = (\x => \y => print_ExprSG show (left (rightDep (semiGroupReduce (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (test5' x y))))) 78 | 79 | 80 | -- new test 81 | 82 | new_a : (x:Nat) -> (y:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x, y] (x + 4) 83 | new_a x y = PlusSG _ _ (VarSG _ _ _ (RealVariable _ _ _ _ FZ)) (ConstSG _ _ _ _ 4) 84 | 85 | newTest_a : (x:Nat) -> (y:Nat) -> (x+4 = x+4) 86 | newTest_a x y = let (Just ok) = semiGroupDecideEq (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (new_a x y) (new_a x y) in ok 87 | -- ok 88 | 89 | 90 | new_b : (x:Nat) -> (y:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x, y] (4 + y) 91 | new_b x y = PlusSG _ _ (ConstSG _ _ _ _ 4) (VarSG _ _ _ (RealVariable _ _ _ _ (FS FZ))) 92 | 93 | newTest_b : (x:Nat) -> (y:Nat) -> (4+y = 4+y) 94 | newTest_b x y = let (Just ok) = semiGroupDecideEq (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (new_b x y) (new_b x y) in ok 95 | -- ok 96 | 97 | 98 | new_c : (x:Nat) -> (y:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x, y] (x + (4 + y)) 99 | new_c x y = PlusSG _ _ (VarSG _ _ _ (RealVariable _ _ _ _ FZ)) 100 | (PlusSG _ _ (ConstSG _ _ _ _ 4) (VarSG _ _ _ (RealVariable _ _ _ _ (FS FZ)))) 101 | 102 | new_d : (x:Nat) -> (y:Nat) -> ExprSG (%instance) (\x => x) (FakeSetAndMult (semiGroup_to_set (%instance))) [x, y] (x + (4 + y)) 103 | new_d x y = PlusSG _ _ (VarSG _ _ _ (RealVariable _ _ _ _ FZ)) 104 | (PlusSG _ _ (ConstSG _ _ _ _ 4) (VarSG _ _ _ (RealVariable _ _ _ _ (FS FZ)))) 105 | 106 | newTest_c_d : (x:Nat) -> (y:Nat) -> (x+(4+y) = x+(4+y)) 107 | newTest_c_d x y = let (Just ok) = semiGroupDecideEq (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (new_c x y) (new_d x y) in ok 108 | -- ok 109 | 110 | 111 | print_test_c : Nat -> Nat -> String 112 | print_test_c = (\x => \y => print_ExprSG show (left (rightDep (semiGroupReduce (%instance) (FakeSetAndNeg (semiGroup_to_set _)) (new_c x y))))) 113 | -- ok, as expected 114 | 115 | 116 | {- 117 | but_they_are_equal : (x:Nat) -> (y:Nat) -> Maybe (left (rightDep (semiGroupReduce (%instance) (test4' x y))) = (left (rightDep (semiGroupReduce (%instance) (test5' x y))))) 118 | but_they_are_equal = \x => \y => exprSG_eq {c=Nat} (%instance) _ _ (left (rightDep (semiGroupReduce (%instance) (test4' x y)))) (left (rightDep (semiGroupReduce (%instance) (test5' x y)))) 119 | -} 120 | 121 | 122 | 123 | 124 | 125 | 126 | 127 | -------------------------------------------------------------------------------- /README: -------------------------------------------------------------------------------- 1 | -- Franck Slama, Edwin Brady 2 | -- University of St Andrews 3 | -- File README 4 | -- Brief description of what's in this repository 5 | --------------------------------------------------- 6 | 7 | The current version of the hierarchy of provers is 0.6 (05/06/2016), and it works with Idris 0.11 8 | 9 | I/ Summary : 10 | --------------------------------------------------------------------- 11 | 12 | This is the implementation of a collection of tactics for Idris, in Idris. These tactics automatically prove equivalences between terms in algebraic structures, with universally quantified variables. 13 | 14 | II/ News : 15 | --------------------------------------------------------------------- 16 | 17 | New with version 0.6 : The collection of tactics compiles with Idris 0.11 (released on March 2016) 18 | ____________________ 19 | 20 | New with version 0.5 : The collection of tactics compiles with Idris 0.9.15.1, and various improvements 21 | ____________________ 22 | 23 | New with version 0.4 : Each prover can now deal with any equivalence relation, instead of just the propositional equality. 24 | ____________________ 25 | 26 | 27 | III/ Current state (05/06/2016) : 28 | --------------------------------------------------------------------- 29 | 30 | All the provers work, but the automatic reflection does not, because of some issues with Idris, that might be addressed in Idris 2. 31 | 32 | 33 | IV/ How to load the tactics : 34 | --------------------------------------------------------------------- 35 | Simply do : 36 | 37 | idris main.idr -p contrib 38 | 39 | That will load and typecheck the 7 tactics that are implemented at the moment + the tests (which includes instances of the corresponding typeclass) for each of them. 40 | 41 | V/ How to use it : 42 | --------------------------------------------------------------------- 43 | 44 | - The automatic reflection mechanism does not work due to some issues with the way reflection works in Idris. Therefore, you still have to encode by hand the two sides of the equality you wish to prove. 45 | 46 | - You can have a look at the test files commutativeGroup_test.idr and group_test.idr (and the other test files) for some detailed examples. 47 | 48 | 49 | VI/ A fairly big example 50 | --------------------------------------------------------------------- 51 | 52 | A fairly big example can be found in the file binary.idr. 53 | Simply run : 54 | 55 | idris ./PotentialApplications/binary.idr -p contrib 56 | 57 | That will load the tactics and that will generate the proof required (in the term goal_aux). 58 | 59 | To see the proof generated, simply ask for the evaluation of : 60 | 61 | \l,c,bit0,bit1,x,x1,v,v1,known => goal_aux l c bit0 bit1 x x1 v v1 known 62 | 63 | Note that because the automatic reflexion does not work, this file is terrible to read because of the encodings that are done by hand 64 | 65 | VII/ About the technique : 66 | --------------------------------------------------------------------- 67 | 68 | 1) Note that we are interested in the proof of correctness (which is precisely the proof of equality/equivalence that we want) and not by the normalised terms, which are just a vector for building the required proofs. 69 | 70 | 2) We follow a "correct by construction" approach : rather than implementing the normalization of terms and afterwards proving that a normalized term and the original term denote the same value, we do it "bit by bit" : we construct this proof at the same time as we normalize the two terms. 71 | The result if that the proof of correctness (which is in what we are interested) is a lot easier to obtain : for each little step of rewriting, the proof is really trivial. 72 | 73 | 3) We work by reflection in the langage itself (Idris), rather than implementing each tactic in the host language (Haskell). That's something nice, especially for a future bootstrapping of Idris. 74 | 75 | 4) Each prover uses the algebraic provers of the structures underneath : we avoid duplication between different levels as much as possible, which implies converting terms between different structures. 76 | For instance, dealing with neutral elements (which appears for the first time at the Monoid level) is NOT re-implemented in the structures above (ie, Group, CommutativeGroup, etc). Instead, the Group prover converts the terms, and calls the Monoid prover. 77 | Some encodings had to be developped for doing some of these conversions (when we convert to a structure which is less expressive than the initial structure, as it is the case when converting from a ring to a commutative group, because the possibility to express products is lost). 78 | 79 | 80 | 81 | -------------------------------------------------------------------------------- /algebraicProvers.ipkg: -------------------------------------------------------------------------------- 1 | package algebraicProvers 2 | 3 | modules = Provers.tools, Provers.mathsResults, Provers.dataTypes, 4 | Provers.magma_reduce, Provers.semiGroup_reduce, Provers.monoid_reduce, 5 | Provers.commutativeMonoid_reduce, Provers.group_reduce, 6 | Provers.commutativeGroup_reduce, Provers.magma_test, 7 | Provers.semiGroup_test, Provers.monoid_test, 8 | Provers.commutativeMonoid_test, Provers.group_test, 9 | Provers.commutativeGroup_test 10 | -------------------------------------------------------------------------------- /autoAssoc/autoAssocNat.idr: -------------------------------------------------------------------------------- 1 | module autoAssocNat 2 | 3 | 4 | import Decidable.Equality 5 | import Data.Fin 6 | import Data.Vect 7 | import autoAssoc_tools 8 | 9 | -- Expr is a reflection of a Nat, indexed over the concrete Nat, 10 | -- and over a set of Nat variables. 11 | 12 | 13 | -- First attempt presented in the paper, without the index 14 | using (x : Nat, y : Nat, G : Vect n Nat) 15 | data XExpr : (G : Vect n Nat) -> Type where 16 | XPlus : XExpr G -> XExpr G -> XExpr G 17 | XVar : (i:Fin n) -> XExpr G 18 | XZero : XExpr G 19 | 20 | -- Second attempt, also presented in the paper 21 | data Expr : (G : Vect n Nat) -> Nat -> Type where 22 | Plus : Expr G x -> Expr G y -> Expr G (x + y) 23 | Var : (i : Fin n) -> Expr G (index i G) 24 | Zero : Expr G Z 25 | 26 | 27 | -- Fully left associative nat expressions 28 | 29 | data LExpr : (G : Vect n Nat) -> Nat -> Type where 30 | LPlus : LExpr G x -> (i : Fin n) -> LExpr G (x + index i G) 31 | LZero : LExpr G Z 32 | 33 | -- Convert an expression to a left associative expression, and return 34 | -- a proof that the rewriting has an equal interpretation to the original 35 | -- expression. 36 | 37 | -- The idea is that we use this proof to build a proof of equality of 38 | -- list Plusends 39 | 40 | expr_l : Expr gam x -> (x' ** (LExpr gam x', x = x')) 41 | expr_l Zero = (_ ** (LZero, Refl)) 42 | expr_l (Var i) = (_ ** (LPlus LZero i, Refl)) 43 | expr_l (Plus ex ey) = 44 | let (x' ** (ex', px)) = expr_l ex in 45 | let (y' ** (ey', py)) = expr_l ey in 46 | let (res ** (normRes, Pres)) = plusLExpr ex' ey' in 47 | (res ** (normRes, rewrite px in (rewrite py in Pres))) 48 | where 49 | plusLExpr : {gam : Vect n Nat} -> {x, y : Nat} 50 | -> LExpr gam x -> LExpr gam y 51 | -> (z ** (LExpr gam z, x+y=z)) 52 | plusLExpr {x=x} ex LZero = 53 | (_ ** (ex, rewrite (plusZeroRightNeutral x) in Refl)) 54 | plusLExpr ex (LPlus e i) = 55 | let (xRec ** (rec, prfRec)) = plusLExpr ex e in 56 | (_ ** (LPlus rec i, ?MplusLExpr)) 57 | 58 | -- ...and back again 59 | 60 | total 61 | l_expr : LExpr G x -> Expr G x 62 | l_expr LZero = Zero 63 | l_expr (LPlus xs i) = Plus (l_expr xs) (Var i) 64 | 65 | -- Convert an expression to some other equivalent expression (which 66 | -- just happens to be normalised to left associative form) 67 | 68 | total 69 | reduce : Expr G x -> (x' ** (Expr G x', x = x')) 70 | reduce e = let (x' ** (e', prf)) = expr_l e in 71 | (x' ** (l_expr e', prf)) 72 | 73 | 74 | -- Build a proof that two expressions are equal. If they are, we'll know 75 | -- that the indices are equal. 76 | 77 | total 78 | eqExpr : (e : Expr G x) -> (e' : Expr G y) -> 79 | Maybe (e = e') 80 | eqExpr (Plus x y) (Plus x' y') with (eqExpr x x', eqExpr y y') 81 | eqExpr (Plus x y) (Plus x y) | (Just Refl, Just Refl) = Just Refl 82 | eqExpr (Plus x y) (Plus x' y') | _ = Nothing 83 | eqExpr (Var i) (Var j) with (decEq i j) 84 | eqExpr (Var i) (Var i) | (Yes Refl) = Just Refl 85 | eqExpr (Var i) (Var j) | _ = Nothing 86 | eqExpr Zero Zero = Just Refl 87 | eqExpr _ _ = Nothing 88 | 89 | total 90 | buildProof : {x : Nat} -> {y : Nat} -> 91 | Expr G ln -> Expr G rn -> 92 | (x = ln) -> (y = rn) -> Maybe (x = y) 93 | buildProof e e' lp rp with (eqExpr e e') 94 | buildProof e e lp rp | Just Refl = ?bp1 95 | buildProof e e' lp rp | Nothing = Nothing 96 | 97 | total 98 | testEq : Expr G x -> Expr G y -> Maybe (x = y) 99 | testEq l r = let (ln ** (l', lPrf)) = reduce l in 100 | let (rn ** (r', rPrf)) = reduce r in 101 | buildProof l' r' lPrf rPrf 102 | 103 | 104 | 105 | -- a couple of test expressions 106 | -- LSH reflected with the first version of the reflected terms 107 | Xe1 : (x, y, z : Nat) -> 108 | XExpr [x, y, z] 109 | Xe1 x y z = XPlus (XPlus (XVar FZ) (XVar (FS FZ))) 110 | (XPlus (XVar FZ) (XVar (FS (FS FZ)))) 111 | 112 | 113 | e1 : (x, y, z : Nat) -> 114 | Expr [x, y, z] ((x + y) + (x + z)) 115 | e1 x y z = Plus (Plus (Var FZ) (Var (FS FZ))) 116 | (Plus (Var FZ) (Var (FS (FS FZ)))) 117 | 118 | e2 : (x, y, z : Nat) -> 119 | Expr [x, y, z] (x + ((y + x) + z)) 120 | e2 xs ys zs = Plus (Var FZ) 121 | (Plus (Plus (Var (FS FZ)) (Var FZ)) (Var (FS (FS FZ)))) 122 | 123 | 124 | -- TODO: need a tactic to run testEq rather than matching on Just ok, 125 | -- since obviously that might fail... 126 | -- At the REPL, try 'magic' to see the generated proof. 127 | 128 | 129 | e1_e2_testEq : (x, y, z : Nat) -> 130 | Maybe (((x + y) + (x + z)) = (x + ((y + x) + z))) 131 | e1_e2_testEq x y z = testEq (e1 x y z) (e2 x y z) 132 | 133 | 134 | magic : (x, y, z : Nat) -> 135 | (((x + y) + (x + z)) = (x + ((y + x) + z))) 136 | magic = \x, y,z => 137 | let (Just ok) = e1_e2_testEq x y z in ok 138 | 139 | 140 | 141 | -- --------------------- 142 | -- Automatic reflection 143 | -- --------------------- 144 | 145 | total 146 | lemmaExtension : {c:Type} -> {n:Nat} -> {m:Nat} -> (g:Vect n c) -> (g':Vect m c) -> (i:Fin n) -> (index i g = index (convertFin _ i m) (g++g')) 147 | lemmaExtension Nil g' (FZ {k=k}) impossible 148 | lemmaExtension (gh::gt) g' (FZ {k=k}) = Refl 149 | lemmaExtension (gh::gt) g' (FS {k=Z} pi) = let proofOfFalse : Void = elimFinZero pi in -- Just elim the element of (Fin 0) that we have in the context to build an inhabitant of False (Void) 150 | ?MlemmaExtension_1 -- Just elim the inhabitant of False that we have constructed in the context 151 | lemmaExtension (gh::gt) g' (FS {k=S pk} pi) = let ih = lemmaExtension gt g' pi in ?MlemmaExtension_2 152 | 153 | 154 | total 155 | weaken : {n:Nat} -> {m:Nat} -> {G:Vect n Nat} -> {x:Nat} -> (G':Vect m Nat) -> (Expr G x) -> (Expr (G ++ G') x) 156 | weaken G' Zero = Zero 157 | weaken G' (Plus e1 e2) = Plus (weaken G' e1) (weaken G' e2) 158 | weaken {n=n} {m=m} {G=G} G' (Var i) = 159 | let pExt = lemmaExtension G G' i in 160 | rewrite pExt in (Var (convertFin _ i m)) 161 | 162 | 163 | total 164 | isElement : {n:Nat} -> (x : a) -> (G : Vect n a) -> Maybe (i:Fin n ** (index i G = x)) 165 | isElement x [] = Nothing 166 | isElement x (y :: ys) with (prim__syntactic_eq _ _ x y) 167 | isElement x (x :: ys) | Just Refl = Just (FZ ** Refl) -- [| Stop |] 168 | isElement x (y :: ys) | Nothing = let recCall = isElement x ys in 169 | -- [| Pop (isElem x ys) |] 170 | case recCall of 171 | Nothing => Nothing 172 | Just (i' ** p') => Just ((FS i') ** ?MisElement_1) 173 | 174 | -- Reflects Nat to Expr 175 | %reflection 176 | total 177 | reflectNat : {n:Nat} -> (G : Vect n Nat) -> (x:Nat) -> (m ** (G' : Vect m Nat ** (Expr (G ++ G') x))) 178 | reflectNat {n=n} G Z = (Z ** ([] ** (Zero {n=n+0} {G=G++[]}))) -- What the hell. Why do I have to give precisely the Z for m ? 179 | reflectNat G (x + y) = 180 | let (_ ** (G' ** x')) = (reflectNat G x) in 181 | let (_ ** (G'' ** y')) = (reflectNat (G ++ G') y) in 182 | let result = Plus (weaken G'' x') y' in 183 | (_ ** ((G' ++ G'') ** ?MreflectNat_1)) -- The proof MreflectNat_1 will use (sym (vectAppendAssociative G G' G'')) 184 | reflectNat {n=n} G t with (isElement t G) 185 | | Just (i ** p) = let result = Var {G=G++[]} (convertFin n i Z) in 186 | (Z ** ([] ** ?MreflectNat_2)) -- We don't add anything 187 | | Nothing ?= (((S Z) ** ((t :: Vect.Nil) ** Var {G=G++[t]} (lastElement' n)))) -- We add the variable 't' to the context, and we can now reference it with (lastElement' n) 188 | -- The proof will use the fact that (index (lastElement' n) (G ++ [t])) = t 189 | -- By using the lemma elemInBigerVect that we've proved in NewAutoAssoc_tools.idr 190 | {- 191 | -- We don't care of Succ here in fact 192 | reflectNat {n=n} G (S x) with (reflectNat G x) 193 | | (m ** (G' ** x')) with (isElement x (G ++ G')) 194 | | Just (i ** proofIndex) = --let prEqual:(Expr (G++G') (S x) = Expr (G++G') (index i (G++G') + 1)) = ?MreflectNat_1 in 195 | let this = Plus (Var i) x' in 196 | (m ** (G' ** ?MreflectNat_2)) -- (G' ** (App (Var i) x'))) 197 | -} 198 | 199 | 200 | ---------- Proofs ---------- 201 | -- This proof is directly done in the definition now 202 | {- 203 | NewAutoAssoc.Mexpr_l1 = proof 204 | intros 205 | rewrite (sym px) 206 | rewrite (sym py) 207 | exact Pres 208 | -} 209 | 210 | -- FIX IDRIS HERE : Why do we have a "gam" and "gam1" in the context ? We should only have ONE contaxt. It's certainly just some elaboration crap... 211 | NewAutoAssocNat.MplusLExpr = proof 212 | intros 213 | rewrite (sym (plusAssociative x1 x2 (index i gam1))) 214 | rewrite prfRec 215 | exact Refl 216 | 217 | NewAutoAssocNat.bp1 = proof { 218 | intros; 219 | refine Just; 220 | rewrite sym lp; 221 | rewrite sym rp; 222 | exact Refl 223 | } 224 | 225 | -- ----------------------------------- 226 | -- Proofs for the automatic reflection 227 | -- ------------------------------------ 228 | -- Just uses the proof of Void that we have constructed in the context 229 | NewAutoAssocNat.MlemmaExtension_1 = proof 230 | intros 231 | exact (void proofOfFalse) 232 | 233 | NewAutoAssocNat.MlemmaExtension_2 = proof 234 | intros 235 | rewrite (sym ih) 236 | rewrite (pre_convertFin_proofIrr pi _ (LTESucc (GTE_plus pk m)) (GTE_plus (S pk) m)) 237 | mrefine Refl 238 | 239 | NewAutoAssocNat.MisElement_1 = proof 240 | intros 241 | exact (elemInBigerVect _ x p' y) 242 | 243 | 244 | 245 | 246 | -------------------------------------------------------------------------------- /main.idr: -------------------------------------------------------------------------------- 1 | -- Edwin Brady, Franck Slama 2 | -- University of St Andrews 3 | 4 | -- Implementation of tactics that prove equivalences in algebraic structures (Rings, Groups, Monoids, etc), written in Idris, for Idris. 5 | -- To prove equivalence over an abstract structure, we normalize both sides of the potential equality and check that these normal forms are syntactically the same. 6 | -- The normalization is implemented following a correct-by-construction approach, enabled by an original type-safe reflection mechanism. 7 | 8 | -- File main.idr 9 | -- Load all the tactics and associated tests files 10 | ------------------------------------------------------------------- 11 | 12 | module Main 13 | 14 | import Provers.tools 15 | import Provers.mathsResults 16 | import Provers.dataTypes 17 | import Provers.magma_reduce 18 | import Provers.semiGroup_reduce 19 | import Provers.monoid_reduce 20 | import Provers.commutativeMonoid_reduce 21 | import Provers.group_reduce 22 | import Provers.commutativeGroup_reduce 23 | import Provers.ring_reduce 24 | 25 | import Provers.magma_test 26 | import Provers.semiGroup_test 27 | import Provers.monoid_test 28 | import Provers.commutativeMonoid_test 29 | import Provers.group_test 30 | import Provers.commutativeGroup_test 31 | -- import Provers.ring_test -- The test file for Ring is not added now (even if it works!) because of the few auxiliary lemmas about * 32 | -- needed for the definition of the instance of the typeclass Ring which aren't proven yet 33 | 34 | import Provers.reflection 35 | 36 | 37 | 38 | main : IO() 39 | main = putStrLn "The collection of tactics for proving equivalences in algebraic structures seems to be ready to prove stuff!" 40 | 41 | 42 | ---------- Proofs ---------- 43 | 44 | 45 | 46 | 47 | 48 | 49 | 50 | 51 | 52 | 53 | 54 | 55 | 56 | 57 | 58 | 59 | 60 | 61 | 62 | 63 | 64 | 65 | 66 | 67 | 68 | 69 | 70 | 71 | 72 | 73 | 74 | 75 | 76 | 77 | 78 | 79 | 80 | 81 | 82 | 83 | 84 | 85 | 86 | 87 | 88 | 89 | 90 | 91 | 92 | 93 | 94 | 95 | 96 | 97 | 98 | 99 | 100 | 101 | 102 | 103 | 104 | 105 | 106 | 107 | 108 | 109 | 110 | 111 | 112 | 113 | 114 | 115 | 116 | 117 | 118 | 119 | 120 | 121 | 122 | 123 | 124 | 125 | 126 | 127 | 128 | 129 | 130 | 131 | 132 | 133 | 134 | 135 | 136 | 137 | 138 | 139 | 140 | 141 | 142 | 143 | 144 | 145 | 146 | 147 | 148 | 149 | 150 | 151 | 152 | 153 | 154 | 155 | 156 | 157 | 158 | 159 | 160 | 161 | 162 | 163 | 164 | 165 | 166 | 167 | 168 | 169 | 170 | 171 | 172 | 173 | 174 | 175 | 176 | 177 | 178 | 179 | 180 | 181 | 182 | 183 | 184 | 185 | 186 | 187 | 188 | 189 | 190 | 191 | 192 | 193 | 194 | 195 | 196 | 197 | 198 | 199 | 200 | 201 | 202 | 203 | 204 | 205 | 206 | 207 | 208 | 209 | 210 | 211 | 212 | 213 | 214 | 215 | 216 | 217 | 218 | 219 | 220 | 221 | 222 | 223 | 224 | 225 | 226 | 227 | 228 | 229 | 230 | 231 | 232 | 233 | 234 | 235 | 236 | 237 | 238 | 239 | 240 | 241 | 242 | 243 | -------------------------------------------------------------------------------- /others/AutoAssoc.idr: -------------------------------------------------------------------------------- 1 | module AutoAssoc 2 | 3 | 4 | import Decidable.Equality 5 | import Data.Fin 6 | import Data.Vect 7 | 8 | -- Expr is a reflection of a list, indexed over the concrete list, 9 | -- and over a set of list variables. 10 | 11 | using (x : List a, y : List a, G : Vect n (List a)) 12 | data Expr : (G : Vect n (List a)) -> List a -> Type where 13 | App : Expr G x -> Expr G y -> Expr G (x ++ y) 14 | Var : (i : Fin n) -> Expr G (index i G) 15 | ENil : Expr G [] 16 | 17 | -- Fully left associative list expressions 18 | 19 | data RExpr : (G : Vect n (List a)) -> List a -> Type where 20 | RApp : RExpr G x -> (i : Fin n) -> RExpr G (x ++ index i G) 21 | RNil : RExpr G [] 22 | 23 | -- Convert an expression to a right associative expression, and return 24 | -- a proof that the rewriting has an equal interpretation to the original 25 | -- expression. 26 | 27 | -- The idea is that we use this proof to build a proof of equality of 28 | -- list appends 29 | 30 | expr_r : Expr G x -> (x' ** (RExpr G x', x = x')) 31 | expr_r ENil = (_ ** (RNil, Refl)) 32 | expr_r (Var i) = (_ ** (RApp RNil i, Refl)) 33 | expr_r (App ex ey) = let (xl ** (xr, xprf)) = expr_r ex in 34 | let (yl ** (yr, yprf)) = expr_r ey in 35 | appRExpr _ _ xr yr xprf yprf 36 | where 37 | appRExpr : (x', y' : List a) -> 38 | RExpr G x -> RExpr G y -> (x' = x) -> (y' = y) -> 39 | (w' ** (RExpr G w', x' ++ y' = w')) 40 | appRExpr x' y' rxs (RApp e i) xprf yprf 41 | = let (xs ** (rec, prf)) = appRExpr _ _ rxs e Refl Refl in 42 | (_ ** (RApp rec i, ?appRExpr1)) 43 | appRExpr x' y' rxs RNil xprf yprf = (_ ** (rxs, ?appRExpr2)) 44 | 45 | -- ...and back again 46 | 47 | r_expr : RExpr G x -> Expr G x 48 | r_expr RNil = ENil 49 | r_expr (RApp xs i) = App (r_expr xs) (Var i) 50 | 51 | -- Convert an expression to some other equivalent expression (which 52 | -- just happens to be normalised to right associative form) 53 | 54 | reduce : Expr G x -> (x' ** (Expr G x', x = x')) 55 | reduce e = let (x' ** (e', prf)) = expr_r e in 56 | (x' ** (r_expr e', prf)) 57 | 58 | -- Build a proof that two expressions are equal. If they are, we'll know 59 | -- that the indices are equal. 60 | 61 | eqExpr : (e : Expr G x) -> (e' : Expr G y) -> 62 | Maybe (e = e') 63 | eqExpr (App x y) (App x' y') with (eqExpr x x', eqExpr y y') 64 | eqExpr (App x y) (App x y) | (Just Refl, Just Refl) = Just Refl 65 | eqExpr (App x y) (App x' y') | _ = Nothing 66 | eqExpr (Var i) (Var j) with (decEq i j) 67 | eqExpr (Var i) (Var i) | (Yes Refl) = Just Refl 68 | eqExpr (Var i) (Var j) | _ = Nothing 69 | eqExpr ENil ENil = Just Refl 70 | eqExpr _ _ = Nothing 71 | 72 | buildProof : {x : List a} -> {y : List a} -> 73 | Expr G ln -> Expr G rn -> 74 | (x = ln) -> (y = rn) -> Maybe (x = y) 75 | buildProof e e' lp rp with (eqExpr e e') 76 | buildProof e e lp rp | Just Refl = ?bp1 77 | buildProof e e' lp rp | Nothing = Nothing 78 | 79 | testEq : Expr G x -> Expr G y -> Maybe (x = y) 80 | testEq l r = let (ln ** (l', lPrf)) = reduce l in 81 | let (rn ** (r', rPrf)) = reduce r in 82 | buildProof l' r' lPrf rPrf 83 | 84 | 85 | -- a couple of test expressions 86 | 87 | e1 : (xs, ys, zs : List a) -> 88 | Expr [xs, ys, zs] ((xs ++ ys) ++ (xs ++ zs)) 89 | e1 xs ys zs = App (App (Var FZ) (Var (FS FZ))) 90 | (App (Var FZ) (Var (FS (FS FZ)))) 91 | 92 | e2 : (xs, ys, zs : List a) -> 93 | Expr [xs, ys, zs] (xs ++ ((ys ++ xs) ++ zs)) 94 | e2 xs ys zs = App (Var FZ) 95 | (App (App (Var (FS FZ)) (Var FZ)) (Var (FS (FS FZ)))) 96 | 97 | 98 | -- TODO: need a tactic to run testEq rather than matching on Just ok, 99 | -- since obviously that might fail... 100 | -- At the REPL, try 'magic {a=Int}' to see the generated proof. 101 | 102 | 103 | e1_e2_testEq : (xs, ys, zs : List a) -> 104 | Maybe (((xs ++ ys) ++ (xs ++ zs)) = (xs ++ ((ys ++ xs) ++ zs))) 105 | e1_e2_testEq xs ys zs = testEq (e1 xs ys zs) (e2 xs ys zs) 106 | 107 | 108 | magic : (xs, ys, zs : List a) -> 109 | (((xs ++ ys) ++ (xs ++ zs)) = (xs ++ ((ys ++ xs) ++ zs))) 110 | magic = \xs, ys,zs => 111 | let (Just ok) = e1_e2_testEq xs ys zs in ok 112 | 113 | 114 | 115 | 116 | 117 | 118 | -- new test 119 | 120 | e3 : (xs, ys : List a) -> 121 | Expr [xs, ys] (([] ++ xs) ++ys) 122 | e3 xs ys = App (App ENil (Var FZ)) (Var (FS FZ)) 123 | 124 | 125 | e4 : (xs, ys : List a) -> 126 | Expr [xs, ys] (xs ++ ys) 127 | e4 xs ys = App (Var FZ) (Var (FS FZ)) 128 | 129 | 130 | e3_e4_testEq : (xs, ys : List a) -> 131 | Maybe ((([] ++ xs) ++ys) = (xs ++ ys)) 132 | e3_e4_testEq xs ys = testEq (e3 xs ys) (e4 xs ys) 133 | 134 | 135 | magic2 : (xs, ys : List a) -> 136 | (([] ++ (xs++ys)) = (xs ++ ys)) 137 | magic2 = \xs, ys => 138 | let (Just ok) = e3_e4_testEq xs ys in ok 139 | 140 | 141 | -- new test 2 142 | 143 | e5 : (xs, ys : List a) -> 144 | Expr [xs, ys] ((xs ++ []) ++ys) 145 | e5 xs ys = App (App (Var FZ) ENil) (Var (FS FZ)) 146 | 147 | 148 | e6 : (xs, ys : List a) -> 149 | Expr [xs, ys] (xs ++ ys) 150 | e6 xs ys = App (Var FZ) (Var (FS FZ)) 151 | 152 | 153 | e5_e6_testEq : (xs, ys : List a) -> 154 | Maybe (((xs ++ []) ++ys) = (xs ++ ys)) 155 | e5_e6_testEq xs ys = testEq (e5 xs ys) (e6 xs ys) 156 | 157 | 158 | magic3 : (xs, ys : List a) -> 159 | (((xs ++ []) ++ys) = (xs ++ ys)) 160 | magic3 = \xs, ys => 161 | let (Just ok) = e5_e6_testEq xs ys in ok 162 | 163 | 164 | 165 | ---------- Proofs ---------- 166 | 167 | AutoAssoc.appRExpr1 = proof { 168 | intros; 169 | rewrite sym xprf; 170 | rewrite sym yprf; 171 | rewrite prf; 172 | mrefine appendAssociative 173 | } 174 | 175 | 176 | appRExpr2 = proof { 177 | intros; 178 | rewrite xprf; 179 | rewrite sym yprf; 180 | rewrite appendNilRightNeutral x'; 181 | exact Refl 182 | } 183 | 184 | bp1 = proof { 185 | intros; 186 | refine Just; 187 | rewrite sym lp; 188 | rewrite sym rp; 189 | exact Refl 190 | } 191 | 192 | 193 | -------------------------------------------------------------------------------- /others/axiomInduction_byRecursivity.idr: -------------------------------------------------------------------------------- 1 | data T : Type where 2 | T1 : T 3 | T2 : T -> T 4 | 5 | {- Induction principle for T -} 6 | total 7 | ind_T : (P:T->Type) -> (P T1) -> ((t:T) -> P t -> P (T2 t)) 8 | -> ((t:T) -> P t) 9 | ind_T P H1 H2 T1 = H1 10 | ind_T P H1 H2 (T2 t) = let ihn = ind_T P H1 H2 t in 11 | H2 t ihn 12 | 13 | 14 | {- So, clearly, we can prove the induction principle for any inductive type, 15 | simple because we can define recursive functions on smaller terms. So either we 16 | consider that recursive definitions is a primitive thing, and thus ind_T can be proven 17 | as we just did, or we take a purely logical point of view (forgetting recursive definitions), 18 | and we consider ind_T to be an axiom -} 19 | 20 | 21 | {- I think in Coq, the induction principle is more primitive, and recursive definitions 22 | are translated automatically by the use of "T_ind" principles that are axioms automatically 23 | added to the theory when we define a new indutive type T. Let's check in axiomInduction_byRecursivity_Coq.v 24 | with the use of "print" for printing the proof term -} 25 | 26 | -------------------------------------------------------------------------------- /others/beqExpr_proof_by_case.v: -------------------------------------------------------------------------------- 1 | Require Import Bool. 2 | 3 | Inductive Expr := 4 | | Const1 : Expr 5 | | Const2 : Expr 6 | | Plus : Expr -> Expr -> Expr. 7 | 8 | Fixpoint Expr_eq (e1:Expr) (e2:Expr) : bool := 9 | match (e1,e2) with 10 | | (Const1, Const1) => true 11 | | (Const2, Const2) => true 12 | | (Plus a1 b1, Plus a2 b2) => (Expr_eq a1 a2) && (Expr_eq b1 b2) 13 | | _ => false 14 | end. 15 | 16 | Fixpoint reify (e:Expr) : nat := 17 | match e with 18 | | Const1 => O 19 | | Const2 => (S O) 20 | | (Plus a b) => (reify a) + (reify b) 21 | end. 22 | 23 | 24 | Lemma Expr_eq_correct : forall (e1 e2:Expr), Expr_eq e1 e2 = true -> reify e1 = reify e2. 25 | Proof. 26 | intros e1 e2 H. 27 | induction e1. 28 | destruct e2. 29 | reflexivity. 30 | inversion H. 31 | inversion H. 32 | destruct e2. 33 | inversion H. 34 | reflexivity. 35 | inversion H. 36 | induction e2. 37 | inversion H. 38 | inversion H. 39 | inversion H. 40 | simpl. 41 | assert (reify e1_1 = reify e2_1). 42 | apply Expr_eq_correct. 43 | 44 | apply IHe2_1. 45 | 46 | 47 | 48 | 49 | 50 | -------------------------------------------------------------------------------- /others/exampleProofByHand.idr: -------------------------------------------------------------------------------- 1 | -- Proof of (l1 ++ l2) ++ (l3 ++l4) = (l1 ++ (l2 ++l3) ++l4) 2 | -- Without automation 3 | 4 | goal : {A:Type} -> (l1:List A) -> (l2:List A) -> (l3:List A) -> (l4:List A) 5 | -> (l1++l2) ++ (l3++l4) = (l1 ++ (l2++l3) ++ l4) 6 | goal l1 l2 l3 l4 = ?MX 7 | 8 | 9 | 10 | ---------- Proofs ---------- 11 | Main.MX = proof 12 | intros 13 | rewrite (appendAssociative l2 l3 l4) -- (l2 + l3) + l4 = (l2 + (l3 + l4)) 14 | rewrite (appendAssociative l1 l2 (l3 ++ l4)) -- (l1 + l2) + (l3 + l4) = l1 + (l2 + (l3 + l4)) 15 | exact refl 16 | 17 | -------------------------------------------------------------------------------- /others/howToDoPerfsTest.txt: -------------------------------------------------------------------------------- 1 | How to do perfs tests : 2 | 3 | In a terminal 4 | --------------- 5 | 6 | idris ./Provers/myBinary_new.idr -p contrib +RTS -k16M -RTS 7 | 8 | 9 | 10 | In a second one 11 | ---------------- 12 | 13 | idris --client ':l ./Provers/myBinary_new.idr -p contrib' 14 | 15 | In order to load the same file 16 | 17 | Then : 18 | 19 | time (idris --client '(\l, c, bit0, bit1, x, x1, v, v1, known => goal_final l c bit0 bit1 x x1 v v1 known)') 20 | 21 | 22 | 23 | 24 | OR : 25 | 26 | (\l, c, bit0, bit1, x, x1, v, v1, known => goalProven l c bit0 bit1 x x1 v v1 known) -------------------------------------------------------------------------------- /others/problemsWithIdris/testReflection2Funs.idr: -------------------------------------------------------------------------------- 1 | {- 2 | Franck Slama - University of St Andrews 3 | file testReflection2Funs.idr 4 | This files correspond to the problem described here : https://github.com/idris-lang/Idris-dev/issues/2843 5 | Last version : 05 June 2016 6 | Compiles with Idris 0.11 (version released at the end of March 2016) 7 | 8 | This file shows two problems with Idris' reflection mechanism : 9 | - 1) It looks like a function defined with %reflection can't call a subfunction doing a subpart of the reflection. 10 | -> that would be useful for having a handler for constants only for example 11 | (in my work, I need to implement an automatic reflection for any type that behaves as a Ring, 12 | but I can't know beforehand what the constants are going to be, so what I want is that the user passes a function dealing only with the reflection of constants) 13 | - 2) Adding a "_" pattern that should not change anything did change the result 14 | -} 15 | 16 | -- Just a (pretty useless) syntax for reflecting natural numbers with additions, variables and constants 17 | data MyNat : Type where 18 | -- Additions 19 | MyPlus : MyNat -> MyNat -> MyNat 20 | -- Constants 21 | MySucc : MyNat -> MyNat 22 | MyZero : MyNat 23 | -- Everything else is seen as a "variable" 24 | AVariable : MyNat -- We forget "which variable it was", but we don't care here for this little example 25 | 26 | -- -------------------------- 27 | -- THAT WORKS : 28 | -- -------------------------- 29 | 30 | %reflection 31 | reflectNatToMyNat : Nat -> MyNat 32 | reflectNatToMyNat (a+b) = MyPlus (reflectNatToMyNat a) (reflectNatToMyNat b) 33 | reflectNatToMyNat (S px) = MySucc (reflectNatToMyNat px) 34 | reflectNatToMyNat Z = MyZero 35 | reflectNatToMyNat _ = AVariable 36 | 37 | -- Here, I get (MyPlus AVariable AVariable) which is what I want : OK 38 | test1 : Nat -> Nat -> MyNat 39 | test1 = (\x,y => reflectNatToMyNat (x+y)) 40 | 41 | 42 | 43 | -- -------------------------- 44 | -- THE FIRST PROBLEM : 45 | -- -------------------------- 46 | 47 | -- This is the type of a handler for constants only 48 | TypeReflectConstants : Type 49 | TypeReflectConstants = Nat -> Maybe(MyNat) 50 | 51 | -- Now, we take as parameter a handler for dealing with the reflection of constants 52 | total 53 | %reflection 54 | reflectNatToMyNat' : (reflectCst : TypeReflectConstants) -> Nat -> MyNat 55 | reflectNatToMyNat' reflectCst (a+b) = MyPlus (reflectNatToMyNat' reflectCst a) (reflectNatToMyNat' reflectCst b) 56 | reflectNatToMyNat' reflectCst t = 57 | case reflectCst t of 58 | -- If the function doing the reflection for constants has decided that t is a constant (and has done the reflection), then we trust it 59 | Just this => this 60 | -- Otherwise, we conclude that we have a variable 61 | Nothing => AVariable 62 | -- _ => AVariable 63 | 64 | -- We define a handler for constants... 65 | total 66 | %reflection 67 | reflectNatConstants : TypeReflectConstants 68 | reflectNatConstants Z = Just MyZero 69 | reflectNatConstants (S px) = case reflectNatConstants px of 70 | Just something => Just (MySucc something) 71 | Nothing => Nothing -- We lose the ability to deal with things like (S x) where x is a variable, but we decide that we don't care 72 | -- This is just the handler for constants, so we simply return Nothing if what we had in input wasn't a constant of Nat 73 | reflectNatConstants _ = Nothing 74 | 75 | 76 | -- I DO NOT get (MyPlus AVariable AVariable) : IT FAILS WITH A "CASE BLOCK IN...". WHY ? 77 | test2 : Nat -> Nat -> MyNat 78 | test2 = (\x,y => reflectNatToMyNat' reflectNatConstants (x+y)) 79 | 80 | -- ------------------------------------------------------------ 81 | -- SECOND PROBLEM WITH reflectNatToMyNat' : 82 | -- ------------------------------------------------------------ 83 | 84 | -- Here I should get (MySucc MyZero), and I do get it 85 | test3 : MyNat 86 | test3 = reflectNatToMyNat' reflectNatConstants (S Z) 87 | -- BUT, if I de-comment the last line "_ => AVariable" of reflectNatToMyNat' (which should be useless), I don't get the same answer (indeed, I get AVariable!) ! Why ??? 88 | 89 | 90 | 91 | 92 | -------------------------------------------------------------------------------- /others/problemsWithIdris/testStackProof.idr: -------------------------------------------------------------------------------- 1 | A : Type 2 | A = ?MA 3 | B : Type 4 | B = ?MB 5 | C : Type 6 | C = ?MC 7 | D : Type 8 | D = ?MD 9 | E : Type 10 | E = ?ME 11 | F : Type 12 | F = ?MF 13 | G : Type 14 | G = ?MG 15 | 16 | th1 : A -> B -> C 17 | th1 = ?M1 18 | 19 | th2 : D -> E -> A 20 | th2 = ?M2 21 | 22 | th3 : F -> G -> B 23 | th3 = ?M3 24 | 25 | ax1 : D 26 | ax1 = ?M4 27 | 28 | ax2 : E 29 | ax2 = ?M5 30 | 31 | ax3 : F 32 | ax3 = ?M6 33 | 34 | ax4 : G 35 | ax4 = ?M7 36 | 37 | 38 | -- Question : After applying th1, we need to produce a proof of both A and B. In order to prove A, we can apply th2. Now, does it asks for a proof of D (expected) or B ? 39 | -- Answer : in Idris 0.9.14 AND 0.9.15 (and earlier versions), we do NOT have the expected behaviour. 40 | my_theorem : C 41 | my_theorem = ?Mgoal -------------------------------------------------------------------------------- /others/problemsWithIdris/unfold_type.idr: -------------------------------------------------------------------------------- 1 | 2 | import Decidable.Equality 3 | import Data.Fin 4 | import Data.Vect 5 | 6 | 7 | %default total 8 | 9 | 10 | -- THE PROBLEM : 11 | -- --------------- 12 | -- we want to write and prove this, but we can't because that doesn't typecheck... 13 | -- index_lemma : {T:Type} -> {n:Nat} -> {n':Nat} -> (i:Fin n) -> (vec : Vect n T) -> (eq:n=n') -> (index i vec = index (replace eq i) vec) 14 | 15 | 16 | -- So my idea is to first define a function which produces the type that we want 17 | -- (and we can write it thanks to a dependent pattern match on the proof of equality eq) : 18 | index_lemma_type : {T:Type} -> {n:Nat} -> {n':Nat} -> (i:Fin n) -> (vec : Vect n T) -> (eq:n=n') -> Type 19 | index_lemma_type i vec eq with (eq) 20 | index_lemma_type i vec eq | (Refl) = (index i vec = index (replace eq i) vec) 21 | 22 | -- And now we can write the type of the function that we want 23 | -- (but we can't define it well...) 24 | index_lemma : {T:Type} -> {n:Nat} -> {n':Nat} -> (i:Fin n) -> (vec : Vect n T) -> (eq:n=n') -> (index_lemma_type i vec eq) 25 | index_lemma i vec eq = ?THERE_IS_A_PROBLEM -- there is a problem because idris doesn't unfold index_lemma_type and seems to be stuck in a "with block in...", 26 | -- like if index_lemma_type would not be total. But it is total. What's the problem ? How to solve it and to define this proof that I need ? 27 | 28 | -- However, it looks like I can prove it if I do again a pattern matching on the proof eq : 29 | index_lemma' : {T:Type} -> {n:Nat} -> {n':Nat} -> (i:Fin n) -> (vec : Vect n T) -> (eq:n=n') -> (index_lemma_type i vec eq) 30 | index_lemma' i vec eq with (eq) 31 | index_lemma i vec eq | (Refl) = Refl -- Here I can do it. But can I use this property now ? Have I effectively prove index i vec = index (replace eq i) vec ? 32 | 33 | -- Just a little tool for a test 34 | plus_one_equals_succ : (n:Nat) -> (n+1 = S n) 35 | plus_one_equals_succ Z = Refl 36 | plus_one_equals_succ (S pn) = let p_ihn : (pn + 1 = S pn) = plus_one_equals_succ pn in ?Mplus_one_equals_succ_1 37 | 38 | 39 | -- Can I use the property index_lemma into another lemma now ? 40 | other_lemma : {T:Type} -> {pn:Nat} -> (i:Fin (S pn)) -> (vec: Vect (S pn) T) -> T 41 | other_lemma {pn=pn} i vec = 42 | let paux = index_lemma' i vec (sym(plus_one_equals_succ pn)) in ?NO_I_CANT_USE_IT 43 | -- No I can't : I don't get any information because Idris doesn't unfold the body of index_lemma_type 44 | 45 | 46 | -- And what if I try this stupid thing ? 47 | -- Well, that's not even accepted... 48 | {- 49 | other_lemma' : {T:Type} -> {pn:Nat} -> (i:Fin (S pn)) -> (vec: Vect (S pn) T) -> T 50 | other_lemma' {pn=pn} i vec with (sym(plus_one_equals_succ pn)) 51 | other_lemma' {pn=pn} i vec | (Refl) = 52 | let paux = index_lemma' i vec (sym(plus_one_equals_succ pn)) in ?NO_I_CANT_USE_IT 53 | -} 54 | 55 | 56 | 57 | 58 | -- Proofs 59 | 60 | Mplus_one_equals_succ_1 = proof 61 | intros 62 | rewrite p_ihn 63 | exact Refl -------------------------------------------------------------------------------- /others/sortedList.idr: -------------------------------------------------------------------------------- 1 | interface Order c where 2 | (>=) : c -> c -> Type 3 | 4 | -- Sorted lists (in decreasing order) by construction (with embedded proofs) 5 | -- The (Maybe T) index is the current maximum, at the head of the list 6 | data SortedList : (T:Type) -> (Order T) -> (Maybe T) -> Type where 7 | Nil : {T:Type} -> {ord:Order T} -> SortedList T ord Nothing 8 | Singleton : {T:Type} -> {ord:Order T} -> (x:T) -> SortedList T ord (Just x) 9 | Cons : {T:Type} -> {ord:Order T} -> {max:T} -> (l:SortedList T ord (Just max)) -> (x:T) -> (pr:x >= max) 10 | -> SortedList T ord (Just x) 11 | 12 | 13 | -- Small example 14 | 15 | implementation [natOrd] Order Nat where 16 | (>=) x y = GTE x y 17 | 18 | -- 2 >= 1 19 | pr1 : GTE 2 1 20 | pr1 = LTESucc LTEZero 21 | 22 | -- 3 >= 2 23 | pr2 : GTE 3 2 24 | pr2 = LTESucc (LTESucc LTEZero) 25 | 26 | -- [3,2,1] 27 | sortedList1 : SortedList Nat Main.natOrd (Just 3) 28 | sortedList1 = Cons (Cons (Singleton 1) 2 pr1) 3 pr2 29 | 30 | -------------------------------------------------------------------------------- /others/test_Ltac_Coq.v: -------------------------------------------------------------------------------- 1 | Ltac test n := 2 | match n with 3 | | O => true 4 | | S _ => O 5 | end. 6 | 7 | (* This is not accepted *) 8 | (* 9 | Definition myfun (n:nat) : nat := 10 | match n with 11 | | O => O 12 | | S pn => test pn 13 | end. 14 | *) 15 | 16 | Ltac test2 n := 17 | apply n. 18 | 19 | (* But this is accepted *) 20 | Definition myfun (n:nat) : nat. 21 | case n. 22 | (try (test2 5)). 23 | (* ok *) 24 | 25 | -------------------------------------------------------------------------------- /others/traditionalApproach.v: -------------------------------------------------------------------------------- 1 | (* Franck Slama 2 | University of St Andrews 3 | traditionalApproach.v 4 | 5 | This file is an axiomatic formalisation of how 6 | our tactic could be developped in a "traditional approach", 7 | ie without our type-safe reflection mechanism. 8 | This code and these proofs conform to my PhD thesis, section 4.8.1 9 | *) 10 | 11 | (* Our universe *) 12 | 13 | Axiom Expr : Type. 14 | Axiom c : Type. 15 | Axiom equiv : c -> c -> Prop. 16 | Notation "x ~= y" := (equiv x y) (at level 75). 17 | Axiom eq_refl : forall x, x ~= x. 18 | Axiom eq_sym : forall x y, x ~= y -> y ~= x. 19 | Axiom eq_trans : forall x y z, (x ~= y) -> (y ~= z) -> (x ~= z). 20 | 21 | Lemma eq_preserves_eq : forall x y c1 c2, (x ~= c1) -> (y ~= c2) -> (c1 ~= c2) -> (x ~= y). 22 | intros. 23 | assert (x ~= c2). 24 | apply eq_trans with (y:=c1). 25 | assumption. 26 | assumption. 27 | assert (c2~=y). 28 | apply eq_sym. 29 | assumption. 30 | apply eq_trans with (y:=c2). 31 | assumption. 32 | assumption. 33 | Qed. 34 | 35 | (* reify, reflect, and their correctness lemma *) 36 | 37 | Axiom reify : Expr -> c. 38 | Axiom reflect : c -> Expr. 39 | 40 | Axiom reflect_and_reify_correct : forall x:c, reify (reflect x) ~= x. 41 | 42 | (* Normalisation function and correctness lemma *) 43 | 44 | Axiom normalise : Expr -> Expr. 45 | 46 | Axiom normalise_correct : forall e:Expr, reify (normalise e) ~= reify e. 47 | 48 | (* Syntactic equality test and correctness lemma *) 49 | 50 | Axiom beq_Expr : Expr -> Expr -> bool. 51 | 52 | Axiom beq_Expr_correct : forall (e1 e2:Expr), beq_Expr e1 e2 = true -> reify e1 ~= reify e2. 53 | 54 | (* Kernel of the tactic and its correctness lemma *) 55 | 56 | Definition decideEq (x:c) (y:c) : bool := 57 | let ex := reflect x in 58 | let ey := reflect y in 59 | beq_Expr (normalise ex) (normalise ey). 60 | 61 | 62 | Theorem decideEq_correct : forall (x y:c), decideEq x y = true -> x ~= y. 63 | Proof. 64 | intros. 65 | 66 | (* 1) "We first need to obtain a proof H0 of (reify (reflect x) ~= reify (reflect y))..." *) 67 | (* -------------------------------------------------------------------------------------- *) 68 | assert (reify (reflect x) ~= reify (reflect y)). 69 | apply eq_preserves_eq with (c1:=reify (normalise (reflect x))) (c2:=reify (normalise (reflect y))). 70 | 71 | (* "... which mostly involves using two times the lemma normalise_correct..." *) 72 | apply eq_sym. 73 | apply normalise_correct. 74 | apply eq_sym. 75 | apply normalise_correct. 76 | 77 | (* "...and to use one time beq_Expr_correct in the (unfolded) hypothesis" *) 78 | unfold decideEq in H. 79 | apply beq_Expr_correct in H. 80 | assumption. 81 | 82 | (* 2) "Then, we can build the desired proof of x ~= y" *) 83 | (* --------------------------------------------------- *) 84 | apply eq_preserves_eq with (c1:=reify (reflect x)) (c2:=reify (reflect y)). 85 | (* "...by using two calls to the reflect_and_reify_correct lemma..." *) 86 | apply eq_sym. 87 | apply reflect_and_reify_correct. 88 | apply eq_sym. 89 | apply reflect_and_reify_correct. 90 | (* ...and the freshly obtained proof $H0$ *) 91 | apply H0. 92 | 93 | Qed. 94 | 95 | 96 | -------------------------------------------------------------------------------- /paper-CICM-17/aliascnt.sty: -------------------------------------------------------------------------------- 1 | %% 2 | %% This is file `aliascnt.sty', 3 | %% generated with the docstrip utility. 4 | %% 5 | %% The original source files were: 6 | %% 7 | %% aliascnt.dtx (with options: `package') 8 | %% 9 | %% This is a generated file. 10 | %% 11 | %% Project: aliascnt 12 | %% Version: 2009/09/08 v1.3 13 | %% 14 | %% Copyright (C) 2006, 2009 by 15 | %% Heiko Oberdiek 16 | %% 17 | %% This work may be distributed and/or modified under the 18 | %% conditions of the LaTeX Project Public License, either 19 | %% version 1.3c of this license or (at your option) any later 20 | %% version. This version of this license is in 21 | %% http://www.latex-project.org/lppl/lppl-1-3c.txt 22 | %% and the latest version of this license is in 23 | %% http://www.latex-project.org/lppl.txt 24 | %% and version 1.3 or later is part of all distributions of 25 | %% LaTeX version 2005/12/01 or later. 26 | %% 27 | %% This work has the LPPL maintenance status "maintained". 28 | %% 29 | %% This Current Maintainer of this work is Heiko Oberdiek. 30 | %% 31 | %% This work consists of the main source file aliascnt.dtx 32 | %% and the derived files 33 | %% aliascnt.sty, aliascnt.pdf, aliascnt.ins, aliascnt.drv. 34 | %% 35 | \NeedsTeXFormat{LaTeX2e} 36 | \ProvidesPackage{aliascnt}% 37 | [2009/09/08 v1.3 Alias counter (HO)]% 38 | \newcommand*{\newaliascnt}[2]{% 39 | \begingroup 40 | \def\AC@glet##1{% 41 | \global\expandafter\let\csname##1#1\expandafter\endcsname 42 | \csname##1#2\endcsname 43 | }% 44 | \@ifundefined{c@#2}{% 45 | \@nocounterr{#2}% 46 | }{% 47 | \expandafter\@ifdefinable\csname c@#1\endcsname{% 48 | \AC@glet{c@}% 49 | \AC@glet{the}% 50 | \AC@glet{theH}% 51 | \AC@glet{p@}% 52 | \expandafter\gdef\csname AC@cnt@#1\endcsname{#2}% 53 | \expandafter\gdef\csname cl@#1\expandafter\endcsname 54 | \expandafter{\csname cl@#2\endcsname}% 55 | }% 56 | }% 57 | \endgroup 58 | } 59 | \newcommand*{\aliascntresetthe}[1]{% 60 | \@ifundefined{AC@cnt@#1}{% 61 | \PackageError{aliascnt}{% 62 | `#1' is not an alias counter% 63 | }\@ehc 64 | }{% 65 | \expandafter\let\csname the#1\expandafter\endcsname 66 | \csname the\csname AC@cnt@#1\endcsname\endcsname 67 | }% 68 | } 69 | \newcommand*{\AC@findrootcnt}[1]{% 70 | \@ifundefined{AC@cnt@#1}{% 71 | #1% 72 | }{% 73 | \expandafter\AC@findrootcnt\csname AC@cnt@#1\endcsname 74 | }% 75 | } 76 | \def\AC@patch#1{% 77 | \expandafter\let\csname AC@org@#1reset\expandafter\endcsname 78 | \csname @#1reset\endcsname 79 | \expandafter\def\csname @#1reset\endcsname##1##2{% 80 | \csname AC@org@#1reset\endcsname{##1}{\AC@findrootcnt{##2}}% 81 | }% 82 | } 83 | \RequirePackage{remreset} 84 | \AC@patch{addto} 85 | \AC@patch{removefrom} 86 | \endinput 87 | %% 88 | %% End of file `aliascnt.sty'. 89 | -------------------------------------------------------------------------------- /paper-CICM-17/content/1-introduction.tex: -------------------------------------------------------------------------------- 1 | \section{Introduction} 2 | 3 | Proofs assistants like Coq~\cite{BertotC04} 4 | and programming languages like 5 | Agda~\cite{norell2007thesis} and 6 | Idris~\cite{brady2013idris} are based on Intensional Type Theories that contain 7 | two notions of equality: propositional equality, that can be 8 | manipulated in the language, and judgemental (or definitional) equality. 9 | Propositional equality corresponds to the mathematical notion: this is 10 | a proposition that can be assumed, negated, proved or disproved. Since in type 11 | theory, propositions are types~\cite{How80}, the proposition that two 12 | elements $x$ and $y$ are equal corresponds to a type. 13 | If $x$ and $y$ are of type $a$, then the type 14 | $Id_a(x, y)$ represents the proposition ``$x$ is equal to $y$". If this type is 15 | inhabited, then $x$ is said to be provably equal to $y$. Thus, $Id$ is a type family (parameterised by the type $a$) indexed over two elements of $a$, giving $\code{Id}\ (a:\code{Type}):a \rightarrow a \rightarrow \code{Type}$. For 16 | convenience, we write $(\code{Id}_a\ x\ y)$ as $(x\ =_a\ y)$. 17 | 18 | Judgemental equality, on the other hand, is a primitive concept of 19 | the type theory. 20 | Whether or not two expressions are judgementally equal is a matter of 21 | evaluating the definitions. For example, if $f\ :\ \mathbb{N}\ \rightarrow\ 22 | \mathbb{N}$ is defined by $f\ x\ :=\ x\ +\ 2$, then $f\ 5$ is 23 | definitionally equal to $7$. Definitional equality entails unfolding 24 | of functions and reductions, until no more reduction can be 25 | performed. We denote the definitional equality by $\equiv$. 26 | 27 | Judgemental equality is included in propositional equality 28 | because what is equal by definition is provably equal. This is 29 | accomplished by giving a constructor for the type $Id(a,a)$ and no 30 | constructor when ``$a$ is not $b$". In these theories, $Id$ is therefore 31 | implemented with the following type with one constructor: 32 | 33 | \begin{lstlisting} 34 | data Id : a $\rightarrow$ a $\rightarrow$ Type where 35 | Refl : (x : a) $\rightarrow$ Id x x 36 | \end{lstlisting} 37 | 38 | The only way for $(\code{Id}_a\ x\ y)$ to be inhabited is therefore that $x$ and $y$ are equal by definition. In this case, the constructor \code{Refl} helps to create a proof of this equality: \code{(Refl\ x)} is precisely the proof 39 | which says that $x=_ax$. 40 | Here, we are using the notation of Idris, where unbound variables 41 | like \texttt{a} in the definition of \texttt{Id} are 42 | \emph{implicitly} quantified, as a concise programming notation. 43 | 44 | The propositional equality does not only contain the judgemental equality, 45 | however, because a principle of induction is associated with each inductive 46 | type. If \code{T} is an inductive type with a constant constructor and a 47 | recursive constructor, ie, \code{T = 1 + T}, defined in Idris as: 48 | 49 | \begin{lstlisting} 50 | data T : Type where 51 | T0 : T 52 | T1 : T $\rightarrow$ T 53 | \end{lstlisting} 54 | 55 | then we have the following induction principle for \code{T}: \\ 56 | $\code{T\_ind} : \forall P:\code{T} \rightarrow \code{Type},\ (P\ \code{T0}) \rightarrow (\forall t:\code{T},\ P\ t \rightarrow P\ (\code{T1}\ t)) \rightarrow (\forall t:\code{T},\ P\ t)$. 57 | 58 | For example, we can prove that $n+0 = n$ for all $n$ by induction on the $Nat\ 59 | n$, even if $n+0 \not\equiv n$ with the usual definition of $+$, recursive on 60 | its first argument. So, the axiom of induction means the type $Id_{a}(x,y)$ 61 | contains not only the canonical form \code{Refl}, but also those added by 62 | inductive principles. There are therefore things which are \emph{provably} 63 | equal, but not \emph{definitionally} equal. Proving equalities is therefore in 64 | these theories something which isn't automatically decidable by the 65 | type-checker in the general case. 66 | 67 | \subsection{Motivating example: Verified Binary Arithmetic} 68 | \label{sect:motivatingExample} 69 | 70 | Proving that one term is equal to another is common in formal verification, and 71 | proof obligations arise naturally in dependently typed programming when 72 | indexing types over values in order to capture some logical properties. 73 | To demonstrate this, 74 | we revisit an example from previous work~\cite{BradyTFP07} which 75 | shows how proof obligations arise when a type is indexed by natural 76 | numbers. Our goal is to implement a verified 77 | library of binary numbers. To ensure functional correctness, 78 | we define the types \code{Bit} and \code{Binary} indexed over the value they represent (expressed as a natural number): 79 | 80 | \begin{lstlisting} 81 | data Bit : Nat $\rightarrow$ Type where 82 | b0 : Bit Z 83 | b1 : Bit (S Z) 84 | 85 | data Binary : (width : Nat) $\rightarrow$ (value : Nat) $\rightarrow$ Type where 86 | zero : Binary Z Z 87 | (#) : Binary w v $\rightarrow$ Bit bit $\rightarrow$ Binary (S w) (bit + 2 * v) 88 | \end{lstlisting} 89 | 90 | We will write a function to add two binary numbers. 91 | To do so, we begin with an auxiliary function, which adds three bits 92 | (the third is a carry bit), and produces the two bits of 93 | the result, where the first is the more significant bit: 94 | 95 | \begin{lstlisting} 96 | addBit : Bit x $\rightarrow$ Bit y $\rightarrow$ Bit c $\rightarrow$ (bX ** (bY ** 97 | (Bit bX, Bit bY, c + x + y = bY + 2 * bX))) 98 | addBit b0 b0 b0 = (_ ** (_ ** (b0, b0, Refl))) 99 | addBit b0 b0 b1 = (_ ** (_ ** (b0, b1, Refl))) 100 | {- ... remaining cases follow the same pattern ... -} 101 | \end{lstlisting} 102 | 103 | The syntax \code{(n ** t)} denotes a \emph{dependent pair}, where the 104 | type of the second argument \texttt{t} can refer to the first argument 105 | \code{n}. So, we can read this type as: ``there exists a number \code{bX}, 106 | and a number \code{bY}, such that we have two bits \code{Bit bX} and 107 | \code{Bit bY} and the sum of the input bits \code{c}, \code{x} and 108 | \code{y} equals \code{bY + 2 * bX}.'' 109 | For example, on the second line, 110 | which corresponds to the computation $0_2 + 0_2 + 1_2 = (01)_2$, 111 | the function produces this bits \code{b0} and \code{b1}, and 112 | a proof that $0 + 0 + 1 = 1 + (2 \times 0)$. 113 | 114 | We then define the function \code{adc} that adds two binary numbers and a carry 115 | bit. This works for two binary numbers with the same number 116 | of bits, and produces a result with one more bit. We would like to write: 117 | 118 | \begin{lstlisting} 119 | adc : Binary w x $\rightarrow$ Binary w y $\rightarrow$ Bit c $\rightarrow$ Binary (S w) (c + x + y) 120 | adc zero zero carry = zero # carry 121 | adc (numx # bX) (numy # bY) carry 122 | = let (vCarry0 ** (vLsb ** (carry0, lsb, _))) 123 | = addBit bX bY carry in 124 | adc numx numy carry0 # lsb 125 | \end{lstlisting} 126 | 127 | Unfortunately, this definition is rejected because the types do 128 | not match for both patterns. 129 | %The result of the first line \code{adc zero zero carry} is expected 130 | %to have the type: 131 | % 132 | %\codeL{Binary\ 1\ ((c\ + 0)\ + 0)} but we provide a 133 | %term of type: 134 | % 135 | %\code{Binary\ 1\ (c\ + (0\ + (0\ + 0)))}. 136 | For the second case, the expected index is: 137 | 138 | \codeL{((c\ + (bit2\ + (v1\ + (v1\ + 0)))) + (bit\ + (v\ + (v\ + 0))))} 139 | \\ 140 | while we're trying to provide a term indexed over: 141 | 142 | \codeL{vLsb\ + (((vCarry0\ + v1) + v)\ + (((vCarry0\ + v1)\ + v)\ + 0))}. 143 | 144 | The definition of \code{adc} we have given would behave correctly, and it has 145 | \emph{provably} the expected type, but it does not have it \emph{immediately} 146 | or \emph{judgementally}: after full reductions the expected and provided types are still 147 | different. 148 | To make the previous 149 | definition acceptable, we need to solve two proof obligations 150 | \code{adc\_lemma\_1} and \code{adc\_lemma\_2} which demand proofs of equality 151 | between the two types. 152 | For example, using a proof script: 153 | 154 | % EB - I took out a step because without the whole proof, all we really 155 | % need to see is some representative examples of rewritings 156 | \begin{lstlisting} 157 | adc_lemma_2 = proof { 158 | intros; 159 | rewrite sym (plusZeroRightNeutral x); 160 | [...] 161 | rewrite (plusAssociative c (plus bit0 (plus v v)) bit1); 162 | rewrite plusCommutative bit1 (plus v v); 163 | [...] 164 | rewrite (plusAssociative (plus (plus x v) v1) (plus x v) v1); 165 | trivial; 166 | } 167 | \end{lstlisting} 168 | 169 | Such proofs consist of a potentially long sequence of 170 | rewriting steps, each using one of the properties: neutral element, 171 | commutativity, associativity. Without some automation, this sequence of 172 | rewritings must be done by the programmer. Not only is this time consuming, 173 | but a small change in the definition may lead to a different proof obligation, 174 | thus invalidating the proof. 175 | A minor change in the datatype, or 176 | the definition of \code{addBit} or \code{adc} will require us 177 | to do a new proof, and thus, without support from the machine, these 178 | proofs 179 | could become the everyday routine in any dependently-typed language. It is worth mentioning that even without using dependent types, these proof obligations for equalities happen very frequently during the formal certification of most applications. 180 | 181 | Our handwritten proof \code{adc\_lemma\_2} uses only the existence of a neutral 182 | element, and the associativity and commutativity of $+$ on \code{Nat}. Thus, 183 | we're rewriting a term by using the properties of a commutative monoid. 184 | With the right choice of combinators~\cite{Carette12} such proofs could be made 185 | much simpler, but we would like a generic prover for commutative monoids to 186 | find a proof automatically. 187 | 188 | \subsection{Our contributions} 189 | 190 | Provers for some algebraic structures have already been implemented for various 191 | proof assistants, including Coq~\cite{Coq2005} and 192 | Agda\footnote{\url{http://wiki.portal.chalmers.se/agda\%5C?n=Libraries.UsingTheRingSolver}}. 193 | In this paper, we describe an implementation\footnote{The 194 | implementation of our hierarchy of tactics can be found online at 195 | \url{https://github.com/FranckS/RingIdris/Provers}} of an automatic prover for 196 | equalities in a \emph{hierarchy} of algebraic structures, including monoids, 197 | groups and rings (all potentially commutative), for the Idris language, 198 | making the following contributions: 199 | 200 | \begin{enumerate} 201 | 202 | \item We present a type-safe reflection mechanism 203 | (section~\ref{sect:typeSafeReflection}), where the reflected terms are indexed 204 | over the concrete terms, providing a direct way to extract proofs and 205 | guaranteeing that the reflected term is a sound representation. 206 | 207 | \item The normalisation procedures are implemented by following a correct by 208 | construction approach (section~\ref{sect:correctByConstruction}), instead of proving the correctness afterwards with 209 | auxiliary lemmas. 210 | % This approach is much more suitable for programming languages like Idris. 211 | 212 | \item We develop a \emph{hierarchy} of tactics where each tactic reuses the 213 | rewriting machinery of the structure from which it inherits. For 214 | example, simplifying neutral elements is implemented only at the monoid level, 215 | and reused at other levels. 216 | It is challenging 217 | to reuse the prover of a less expressive structure; for 218 | example, using the monoid prover to build the group prover is 219 | tricky because 220 | we lose the ability to express negations ($-x$) and 221 | subtractions ($x-y$). We present encodings 222 | (section~\ref{sect:reusabilityOfTheProvers}) to overcome this problem. 223 | 224 | \end{enumerate} 225 | 226 | The principal novelty is in using 227 | \emph{type-safe reflection}. Working by reflection for implementing tactics 228 | has been done several times, including the implementation of a ring solver 229 | for Coq, but without the type-safety and 230 | correctness-by-construction. We compare our approach with other 231 | implementations in section~\ref{sect:relatedWork}. 232 | 233 | 234 | 235 | -------------------------------------------------------------------------------- /paper-CICM-17/content/2-basicTactic.tex: -------------------------------------------------------------------------------- 1 | \section{A simple proof automation for universally quantified natural numbers and their associative addition} 2 | \label{sect:aSimpleProofAutomation} 3 | 4 | 5 | We will first present the basic ideas of our technique on a simplified problem, in which we aim to deal with universally quantified natural numbers and additions. We would like to be able to automatically generate proofs of this kind of goals : 6 | \\ 7 | \textbf{Example 1.} $\forall (x\ y\ z:Nat),\ (x + y) + (x + z) = x + ((y + x) + z)$ \\ 8 | \ 9 | In this section, we will only work with the properties of associativity and of neutral element for the $+$ operation on $Nat$. 10 | 11 | \subsection{Type-safe reflection} 12 | 13 | 14 | When trying to prove this kind of equalities, the variables are abstracted, and 15 | they become part of the context. The general idea --that will also apply for 16 | the more general problem detailed in the next section-- will be to normalize 17 | both sides of the "potential equality" $x=y$, and afterwards to compare them 18 | with a syntactic equality test. The problem is that such a normalisation 19 | function will need to do different treatments for a ``variable natural number" 20 | (ie, a number which has been universally quantified), for the constant $Z$ and 21 | for a sum. For this reason, we will need to work by reflection, with a datatype 22 | encoding natural numbers composed of ``variable numbers", $Z$, and additions of 23 | these things. The novelty is that we use a a type-safe reflection mechanism 24 | where we index the reflected terms by the concrete value that they represent: 25 | 26 | \begin{lstlisting} 27 | using (x : Nat, y : Nat, $\Gamma$ : Vect n Nat) 28 | data Expr : (Vect n Nat) $\rightarrow$ Nat $\rightarrow$ Type where 29 | Plus : Expr $\Gamma$ x $\rightarrow$ Expr $\Gamma$ y $\rightarrow$ Expr $\Gamma$ (x + y) 30 | Var : (i : Fin n) $\rightarrow$ Expr $\Gamma$ (index i $\Gamma$) 31 | Zero : Expr $\Gamma$ Z 32 | \end{lstlisting} 33 | 34 | For an expression $e_x\ :\ Expr\ \Gamma\ x$, we will say that $e_x$ denotes (or 35 | encodes) the number $x$ in the context $\Gamma$. When an expression is a 36 | variable $Var\ i$, the denoted number is simply the corresponding variable in 37 | the context, ie, $(index\ i\ \Gamma)$. Also, the $Zero$ expression denotes the 38 | natural number $Z$. Finally, if $e_x$ is an expression encoding the number $x$, 39 | and $e_y$ is an expression encoding the number $y$, then the expression $Plus\ 40 | e_x\ e_y$ denotes the concrete natural number $(x + y)$. 41 | 42 | Note that variables are represented using a De Brujin-like index: $Var\:FZ$ 43 | denotes the first variable abstracted, $Var\:(FS\:FZ)$ the second one, and so 44 | on. 45 | 46 | We can define the reflection of the left-hand side of the example 1 like this : 47 | \begin{lstlisting} 48 | e1 : (x, y, z : Nat) $\rightarrow$ Expr [x, y, z] ((x + y) + (x + z)) 49 | e1 x y z = Plus (Plus (Var FZ) (Var (FS FZ))) 50 | (Plus (Var FZ) (Var (FS (FS FZ)))) 51 | \end{lstlisting} 52 | 53 | The two important things to remember with this type-safe reflection is that the reflection of a term is guaranteed to be a faithful encoding of it because of the index, and that this index will help us to build the construction of the desired proof of equality, as we will show in the next subsection. 54 | 55 | 56 | \subsection{A correct by construction approach} 57 | \label{sect:aCorrectByConstructionApproach} 58 | 59 | We want to write the reduction function on a \emph{correct by construction} way, which means that no additional proof should be required after the definition of the function. Thus, $reduce$ will produce the proof that the new Expr freshly produced has the same interpretation as the original Expr, and this will be made easier by the fact that the datatype Expr is indexed over the real --concrete-- natural number : a term of type $Expr\ \Gamma\ x$ is the encoding of the number $x$. 60 | Thus, we can write the type of $reduce$ like this : \\ 61 | $reduce\ :\ Expr\ \Gamma\ x\ \rightarrow\ (x'\ **\ (Expr\ \Gamma\ x',\ x\ =\ x'))$ \\ 62 | The function $reduce$ produces a dependent pair : the new concrete number $x'$, and a pair made of an $Expr\ \Gamma\ x'$ which is the new encoded term indexed over the new concrete number we have just produced, and a proof that old and new --concrete-- natural numbers are equal. 63 | 64 | This function doesn't simply produce an $Expr\ \Gamma\ x$ (with an internal conversion of type) because the proof of $x=x'$ is an essential component for building the desired proof of $x=y$, which is the main problem that we are trying to solve. More precisely, the tactic will work as follow. 65 | We have an expression $e_x$ encoding $x$, and an expression $e_y$ encoding $y$. We will normalize $e_x$ and this will give a new concrete number $x'$, a new expression $e_{x'}:Expr\ \Gamma\ x'$, and a proof of $x=x'$. We will do the same with $e_y$ and we will get a new concrete number $y'$, an expression $e_{y'}:Expr\ \Gamma\ y'$, and a proof of $y=y'$. \\ 66 | It is now enough to simply compare $e_{x'}$ and $e_{y'}$ using a standard syntactic equality test because these two expressions are now supposed to be in normal form. 67 | 68 | 69 | \begin{lstlisting} 70 | eqExpr : (e : Expr $\Gamma$ x) $\rightarrow$ (e' : Expr $\Gamma$ y) $\rightarrow$ Maybe (e = e') 71 | \end{lstlisting} 72 | 73 | 74 | Now, if the two normalised expressions $e_{x'}$ and $e_{y'}$ are equal, then they necessary have the same type\footnote{We are working with the heterogeneous equality JMeq by default in Idris, but as always, the only way to have a proof of a:A = b:B is when A=B}, and therefore $x'=y'$. 75 | By rewriting the two equalities $x=x'$ and $y=y'$ (that we obtained during the normalisations) in the new equality $x'=y'$, we can get a proof of $x=y$. This is what the function $buildProof$ is doing. 76 | 77 | \begin{lstlisting} 78 | buildProof : {x : Nat} $\rightarrow$ {y : Nat} $\rightarrow$ Expr $\Gamma$ x' $\rightarrow$ Expr $\Gamma$ y' 79 | $\rightarrow$ (x = x') $\rightarrow$ (y = y') $\rightarrow$ Maybe (x = y) 80 | buildProof ex' ey' lp rp with (eqExpr ex' ey') 81 | buildProof ex' ex' lp rp | Just Refl = ?MbuildProof 82 | buildProof ex' ey' lp rp | Nothing = Nothing 83 | \end{lstlisting} 84 | 85 | 86 | The argument of type $Expr\ \Gamma\ x'$ is the normalised reflected left hand size of the equality, which represents the value $x'$. Before the normalisation, the reflected LHS was reflecting the value $x$. The same applies to the right-hand side. The function also expects a proof of $x=x'$ and of $y=y'$, that we will be able to provide because the normalisation function has produced them. 87 | 88 | As mentioned, the proof-term that fills the hole represented by the metavariable $MbuildProof$ is just a rewriting of the two equalities that we have : 89 | 90 | \begin{lstlisting} 91 | MbuildProof = proof { 92 | intros; refine Just; rewrite sym p1; rewrite sym p2; exact Refl; 93 | } 94 | \end{lstlisting} 95 | 96 | 97 | Finally, the main function which tries to prove the equality $x=y$ simply has to reduce the two reflected terms encoding the left and the right hand side, and to use the function $buildProof$ in order to compose the two proofs that have been obtained : 98 | 99 | 100 | \begin{lstlisting} 101 | testEq : Expr $\Gamma$ x $\rightarrow$ Expr $\Gamma$ y $\rightarrow$ Maybe (x = y) 102 | testEq ex ey = 103 | let (x' ** (ex', px)) = reduce ex in 104 | let (y' ** (ey', py)) = reduce ey in 105 | buildProof ex' ey' px py 106 | \end{lstlisting} 107 | 108 | 109 | The remaining part is to define the function $reduce$. We decide that the left associative form will be the canonical representation. To help for this task, we define a new datatype which captures the property of being left-associative by its shape : 110 | 111 | \begin{lstlisting} 112 | data LExpr : ($\Gamma$ : Vect n Nat) $\rightarrow$ Nat $\rightarrow$ Type where 113 | LPlus : LExpr $\Gamma$ x $\rightarrow$ (i : Fin n) 114 | $\rightarrow$ LExpr $\Gamma$ (x + index i $\Gamma$) 115 | LZero : LExpr $\Gamma$ Z 116 | \end{lstlisting} 117 | 118 | 119 | This datatype has only two constructors. In fact, it combines the previous $Var$ and $Plus$ constructors so that it becomes impossible to write an expression which isn't left associative, because the constructor LPlus is only recursive on its first argument. 120 | 121 | As part of the normalization, we write a function $expr\_l$ which converts an $Expr\ \Gamma\ x$ to a $LExpr\ \Gamma\ x'$ and which produces a proof of $x=x'$. This function will therefore use the two available properties multiple times, while rewriting the term until the fully left associative desired form is obtained. 122 | 123 | 124 | \begin{lstlisting} 125 | expr_l : Expr $\Gamma$ x $\rightarrow$ (x' ** (LExpr $\Gamma$ x', x = x')) 126 | expr_l Zero = (_ ** (LZero, Refl)) 127 | expr_l (Var i) = (_ ** (LPlus LZero i, Refl)) 128 | expr_l (Plus ex ey) = 129 | let (x' ** (ex', px)) = expr_l ex in 130 | let (y' ** (ey', py)) = expr_l ey in 131 | let (res ** (normRes, Pres)) = plusLExpr ex' ey' in 132 | (res ** (normRes, rewrite px in (rewrite py in Pres))) where 133 | plusLExpr : {$\Gamma$ : Vect n Nat} $\rightarrow$ {x, y : Nat} $\rightarrow$ LExpr $\Gamma$ x 134 | $\rightarrow$ LExpr $\Gamma$ y $\rightarrow$ (z ** (LExpr $\Gamma$ z, x+y=z)) 135 | plusLExpr {x=x} ex LZero = 136 | (_ ** (ex, rewrite (plusZeroRightNeutral x) in Refl)) 137 | plusLExpr ex (LPlus e i) = 138 | let (xRec ** (rec, prfRec)) = plusLExpr ex e in 139 | (_ ** (LPlus rec i, ?MplusLExpr)) 140 | 141 | \end{lstlisting} 142 | 143 | In the case of an addition $Plus\ ex\ ey$, the function $expr\_l$ does the job of normalisation recursively on $ex$ and on $ey$, and then it uses the sub-function $plusLExpr$ to normalise the addition of these two --already normalised-- terms. This sub-function $plusLExpr$ has two kind of simplifications to do. When the second argument is an $LZero$, it simply returns its first arguments along with the justification for this rewriting, which obviously uses $plusZeroRightNeutral$. However, when the second argument is an $LPlus\ e\ i$, it continues recursively by computing $plusLExpr\ ex\ e$, and it finally adds $i$ to it. That had for effect to move the parenthesis on the left, and the correctness of this treatment is going to be justified by the use of $plusAssociative$ in the proof that corresponds to the meta variable $MplusLExpr$ . 144 | 145 | This metavariable $MplusLExpr$ --that expressed the equality between the old and the new index-- requires us to prove the goal : $x1 + (x2 + index\ i\ \Gamma) = xrec + index\ i\ \Gamma$ in a context where we've got, amongst other things, $prfRec : x1 + x2 = xrec$. 146 | By using the property of associativity on the goal, we now need to prove $(x1 + x2) + index\ i\ \Gamma$, which can be done by rewriting the proof $prfRec$ obtained recursively. 147 | 148 | 149 | \begin{lstlisting} 150 | MplusLExpr = proof { 151 | intros 152 | rewrite (sym (plusAssociative x1 x2 (index i $\Gamma$))); 153 | rewrite prfRec; 154 | exact Refl; 155 | } 156 | \end{lstlisting} 157 | 158 | 159 | It is really important to understand that the heart of the automatic construction of the desired proof of $x=y$ happens precisely in the repeated use of $plusZeroRightNeutral$ and $plusAssociative$ that we've seen in the definition of the fixpoint computed by $expr\_l$ and of its meta-variable $MplusLExpr$. The complete proof is produced during the recursive calls to $expr\_l$ that will compose the use of $plusAssociative$ for producing automatically a proof that can replace the arithmetical proofs that we were doing previously by hand with the lemma $adc\_lemma\_2$. 160 | 161 | The normalisation function is just the composition of the functions $expr\_l$ and $l\_expr$, where the later is the conversion back to the type $Expr$ : 162 | 163 | \begin{lstlisting} 164 | reduce : Expr $\Gamma$ x $\rightarrow$ (x' ** (Expr $\Gamma$ x', x = x')) 165 | reduce e = let (x' ** (e', prf)) = expr_l e 166 | in (x' ** (l_expr e', prf)) 167 | \end{lstlisting} 168 | 169 | At the moment, what we've got is not exactly a real tactic, in the sense that we only have a function which produces a value of type $Maybe (x = y)$. A real tactic would be a wrapper of this function that would fail properly with an error message when the two terms are not equal. However, here, when $x\ne y$, the function $testEq$ will simply produce the value $Nothing$. \\ 170 | 171 | 172 | \subsection{Usage of the ``tactic"} 173 | 174 | It's now time to see how to use this minimalist ``tactic". 175 | Let's go back to the example 1. We had defined the expression $e1$ representing the value $((x + y) + (x + z))$ and we now have to produce the encoding for $(x + ((y + x) + z))$, still in the context $[x, y, z]$ of three abstracted variables. 176 | 177 | 178 | \begin{lstlisting} 179 | e2 : (x, y, z : Nat) 180 | $\rightarrow$ Expr [x, y, z] (x + ((y + x) + z)) 181 | e2 x y z = Plus (Var FZ) 182 | (Plus (Plus (Var (FS FZ)) 183 | (Var FZ)) 184 | (Var (FS (FS FZ)))) 185 | \end{lstlisting} 186 | 187 | 188 | The numbers denoted by the expressions $e1$ and $e2$ are equal, and we can generate a proof of this fact by using $testEq$. 189 | 190 | 191 | \begin{lstlisting} 192 | e1_e2_testEq : (x, y, z : Nat) 193 | $\rightarrow$ Maybe (((x + y) + (x + z)) = (x + ((y + x) + z))) 194 | e1_e2_testEq x y z = testEq (e1 x y z) (e2 x y z) 195 | \end{lstlisting} 196 | 197 | 198 | 199 | And if we ask for the evaluation of this term, we obtain $Just$ and a proof of equality between the two underlying concrete values : 200 | 201 | 202 | \begin{lstlisting} 203 | #\x => \y => \z => e1_e2_testEq x y z 204 | 205 | \x => \y => \z => Just (replace (sym (replace (sym (replace (sym 206 | (plusAssociative x 0 y)) (replace (replace (sym (plusZeroRightNeutral 207 | x)) Refl) Refl))) (replace (sym (replace (sym (plusAssociative x 0 z)) 208 | (replace (replace (sym (plusZeroRightNeutral x)) Refl) Refl))) 209 | (replace (sym (plusAssociative (x+y) x z)) [...] 210 | : (x : Nat) $\rightarrow$ (y : Nat) $\rightarrow$ (z : Nat) 211 | $\rightarrow$ Maybe ((x + y) + (x + z) = x + ((y + x) + z)) 212 | \end{lstlisting} 213 | 214 | As expected, this proof uses the properties of associativity ($plusAssociative$) and the property of neutrality of $Z$ for $+$ ($plusZeroRightNeutral$). 215 | 216 | 217 | -------------------------------------------------------------------------------- /paper-CICM-17/content/4-relatedWork.tex: -------------------------------------------------------------------------------- 1 | \section{Related work} 2 | \label{sect:relatedWork} 3 | 4 | Coq's ring solver~\cite{Coq2005}, like ours, is 5 | implemented using proof-by-reflection techniques, but without the guarantees obtained with 6 | our type-safe reflection mechanism, and without the correct-by-construction 7 | approach: first, they define the normalisation of terms, then 8 | they prove correctness of the normalisation with an external lemma: 9 | $\forall\ (e1\ e2\ :\ \code{Expr}),\ beq_{Expr}\ (\code{norm}\ e1)\ 10 | (\code{norm}\ e2)\ =\ true\ \rightarrow\ \code{reify}\ e1\ \simeq\ 11 | \code{reify}\ e2$. 12 | This needs a \code{reify} function, which we do not need. 13 | Furthermore, Coq's prover deals 14 | with rings and semi-rings (commutative or not), but not with any of the 15 | intermediate structures (semi-group, monoid, group, etc). 16 | However, 17 | their implementation has better performances due to the use of sparse normal 18 | form and more optimised algorithms. Our automatic reflection was written with 19 | Idris reflection mechanism which allows pattern matching on syntax, and 20 | their automatic reflection is programmed in Ltac~\cite{DelahayeLTac}, a proof 21 | dedicated and untyped meta-language for the writing of automations. 22 | More recently, the Mtac extension~\cite{Ziliani13} provides a 23 | typed language for implementating proof automation. 24 | 25 | As well as the ring solver, Coq also provides the 26 | Omega solver~\cite{Cregut04}, which solves a goal in Presburger arithmetic 27 | (i.e. a universally quantified formula made of equations and inequations), and 28 | a field~\cite{DelahayeField} decision procedure for real numbers, which plugs 29 | to Coq's ring prover after simplification of the multiplicative inverses. 30 | Agda's reflection 31 | mechanism\footnote{\url{http://wiki.portal.chalmers.se/agda/pmwiki.php?n=ReferenceManual.Reflection}} 32 | gives access to a representation of the current goal (that is, 33 | the required type) at a particular point in a program. 34 | This allows various proof automations to be done in 35 | Agda~\cite{DBLP:conf/mpc/KokkeS15,Lindblad04}. 36 | 37 | Proofs by reflection has been intensively 38 | studied~\cite{ChlipalaBook,Malecha14}, but without anything similar to 39 | the type-safe reflection that we have presented here. 40 | % 41 | If we leave the ground of nice mathematical structures, one can decide 42 | to work with arbitrary rewriting rules, but in the general case there isn't a 43 | complete decision procedure for such systems, because there is usually no 44 | normal form. This is where deduction modulo~\cite{Dowek03,DelahayeModulo} and 45 | proof search heuristics start. 46 | -------------------------------------------------------------------------------- /paper-CICM-17/content/5-conclusions.tex: -------------------------------------------------------------------------------- 1 | \section{Results, conclusions and future work} 2 | 3 | We have implemented a generic solution to the initial problem of index 4 | mismatch (section~\ref{sect:motivatingExample}) when 5 | using indexed types. This solution takes the form of a hierarchy of provers for 6 | equivalences in algebraic structures. These provers are generic in several 7 | ways: they work for many algebraic structures (semi-group, monoid, commutative 8 | monoid, group, commutative group, ring and semi-ring); for any type that 9 | behaves as one of these structures; and, for any equivalence relation on this 10 | type. The implementation is 11 | modular and each prover reuses the prover of the structure from which it 12 | inherits. These provers can automatically prove equivalences between terms, 13 | so the user need not prove obligations by hand, 14 | like \code{adc\_lemma\_2}. Thus, these provers enable the user to focus on the 15 | interesting proofs that requires specific knowledge and creativity, instead of 16 | routine, automatable, lemmas. 17 | 18 | Our correct-by-construction method involved the design of a type-safe 19 | reflection mechanism where reflected terms are indexed over concrete inputs, 20 | and from which we are able to extract proofs directly. Unlike Coq's and Agda's 21 | ring provers, we do not have the duplication that arise when separating the 22 | computational content from the proof of correctness. Instead, construction of 23 | the proof is done step by step, following the construction of the normalised 24 | terms. In addition to avoiding redundancy, this simplifies the proof 25 | generation considerably. This development shows that if dependent types 26 | effectively bring some new problems, they are also very expressive tools for 27 | building correct-by-construction software where the development is \emph{driven 28 | by the types}. 29 | 30 | This work can be extended to build new provers for less common algebraic 31 | structures and for more specific structures. For example, regular expressions 32 | form a ``pre semi-ring'' with some extra axioms. We will refactor the semi-ring 33 | level with the creation of the intermediate structure of pre semi-ring, that 34 | will not necessary have the property that $0$ is an annihilator element for the 35 | product. Then, we could build a specific prover for regular expressions, that 36 | would use the normalisation function of the pre semi-ring level and that would 37 | only have to deal with the specific properties of regular expressions : the 38 | neutral element $\varnothing$ for the concatenation of languages is also a 39 | neutral element for the product of languages, the idempotence of the addition 40 | of languages, and the rules of simplifications for the new Kleene star 41 | operation. 42 | 43 | 44 | -------------------------------------------------------------------------------- /paper-CICM-17/dtp.bib: -------------------------------------------------------------------------------- 1 | @inproceedings{Carette12, 2 | author = {Jacques Carette and 3 | Russell O'Connor}, 4 | title = {Theory Presentation Combinators}, 5 | booktitle = {International Conference on Intelligent Computer Mathematics (CICM) 2012}, 6 | year = {2012} 7 | } 8 | @inproceedings{BradyTFP07, 9 | author = {Brady, Edwin}, 10 | booktitle = {Trends in Functional Programming ({TFP}'07)}, 11 | title = {{Constructing Correct Circuits: Verification of Functional Aspects of Hardware Specifications with Dependent Types}}, 12 | year = {2007} 13 | } 14 | @inproceedings{Farmer13, 15 | author = {William M. Farmer}, 16 | title = {The Formalization of Syntax-Based Mathematical Algorithms Using Quotation 17 | and Evaluation}, 18 | booktitle = {International Conference on Intelligent Computer Mathematics (CICM) 2013}, 19 | year = {2013} 20 | } 21 | 22 | @inproceedings{DBLP:conf/plpv/Brady13, 23 | author = {Edwin Brady}, 24 | title = {Idris: general purpose programming with dependent types}, 25 | booktitle = {Proceedings of the 7th Workshop on Programming languages meets program 26 | verification, {PLPV} 2013}, 27 | pages = {1--2}, 28 | year = {2013}, 29 | crossref = {DBLP:conf/plpv/2013}, 30 | doi = {10.1145/2428116.2428118}, 31 | timestamp = {Fri, 27 Sep 2013 20:12:08 +0200}, 32 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/plpv/Brady13}, 33 | bibsource = {dblp computer science bibliography, http://dblp.org} 34 | } 35 | 36 | @proceedings{DBLP:conf/plpv/2013, 37 | editor = {Matthew Might and 38 | David Van Horn and 39 | Andreas Abel and 40 | Tim Sheard}, 41 | title = {Proceedings of the 7th Workshop on Programming languages meets program 42 | verification, {PLPV} 2013, Rome, Italy, January 22, 2013}, 43 | publisher = {{ACM}}, 44 | year = {2013}, 45 | url = {http://dl.acm.org/citation.cfm?id=2428116}, 46 | isbn = {978-1-4503-1860-0}, 47 | timestamp = {Fri, 27 Sep 2013 20:12:08 +0200}, 48 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/plpv/2013}, 49 | bibsource = {dblp computer science bibliography, http://dblp.org} 50 | } 51 | 52 | 53 | 54 | @inproceedings{Coq2005, 55 | author = {Gregoire, Benjamin and Mahboubi, Assia}, 56 | booktitle = {{Theorem Proving in Higher Order Logics (TPHOLS 2005)}}, 57 | file = {:Users/edwin/Documents/Mendeley Desktop/Gregoire, Mahboubi/Proving Equalities in a Commutative Ring Done Right in Coq/Gregoire, Mahboubi - Proving Equalities in a Commutative Ring Done Right in Coq - 2005.pdf:pdf}, 58 | pages = {98----113}, 59 | title = {{Proving Equalities in a Commutative Ring Done Right in Coq}}, 60 | year = {2005} 61 | } 62 | 63 | @article{brady2013idris, 64 | author = {Edwin Brady}, 65 | title = {Idris, a general-purpose dependently typed programming language: Design and implementation}, 66 | journal = {Journal of Functional Programming}, 67 | volume = {23}, 68 | issue = {05}, 69 | month = {September}, 70 | year = {2013}, 71 | issn = {1469-7653}, 72 | pages = {552--593}, 73 | numpages = {42}, 74 | doi = {10.1017/S095679681300018X} 75 | } 76 | @inproceedings{How80, 77 | author = {W.A. Howard}, 78 | year = 1980, 79 | title = {The formulae-as-types notion of construction}, 80 | booktitle = {{To H. B. Curry: Essays on Combinatory Logic, Lambda Calculus, and Formalism}}, 81 | editor = {J.R. Seldin and J.P. Hindley} 82 | } 83 | 84 | @article{Dowek03, 85 | author = {Gilles Dowek and 86 | Th{\'{e}}r{\`{e}}se Hardin and 87 | Claude Kirchner}, 88 | title = {Theorem Proving Modulo}, 89 | journal = {J. Autom. Reasoning}, 90 | volume = {31}, 91 | number = {1}, 92 | pages = {33--72}, 93 | year = {2003}, 94 | url = {http://dx.doi.org/10.1023/A:1027357912519}, 95 | doi = {10.1023/A:1027357912519}, 96 | timestamp = {Wed, 04 Feb 2004 12:32:54 +0100}, 97 | biburl = {http://dblp.uni-trier.de/rec/bib/journals/jar/DowekHK03}, 98 | bibsource = {dblp computer science bibliography, http://dblp.org} 99 | } 100 | 101 | @book{BertotC04, 102 | author = {Yves Bertot and 103 | Pierre Cast{\'{e}}ran}, 104 | title = {Interactive Theorem Proving and Program Development - Coq'Art: The 105 | Calculus of Inductive Constructions}, 106 | series = {Texts in Theoretical Computer Science. An {EATCS} Series}, 107 | publisher = {Springer}, 108 | year = {2004}, 109 | doi = {10.1007/978-3-662-07964-5}, 110 | isbn = {978-3-642-05880-6}, 111 | timestamp = {Wed, 09 Sep 2015 16:46:05 +0200}, 112 | biburl = {http://dblp.uni-trier.de/rec/bib/series/txtcs/BertotC04}, 113 | bibsource = {dblp computer science bibliography, http://dblp.org} 114 | } 115 | 116 | @inproceedings{Ziliani13, 117 | author = {Beta Ziliani and 118 | Derek Dreyer and 119 | Neelakantan R. Krishnaswami and 120 | Aleksandar Nanevski and 121 | Viktor Vafeiadis}, 122 | title = {Mtac: a monad for typed tactic programming in Coq}, 123 | booktitle = {{ACM} {SIGPLAN} International Conference on Functional Programming, 124 | ICFP'13}, 125 | pages = {87--100}, 126 | year = {2013}, 127 | doi = {10.1145/2500365.2500579}, 128 | timestamp = {Thu, 19 Sep 2013 11:26:38 +0200}, 129 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/icfp/ZilianiDKNV13}, 130 | bibsource = {dblp computer science bibliography, http://dblp.org} 131 | } 132 | @proceedings{DBLP:conf/icfp/2013, 133 | editor = {Greg Morrisett and 134 | Tarmo Uustalu}, 135 | title = {{ACM} {SIGPLAN} International Conference on Functional Programming, 136 | ICFP'13, Boston, MA, {USA} - September 25 - 27, 2013}, 137 | publisher = {{ACM}}, 138 | year = {2013}, 139 | url = {http://dl.acm.org/citation.cfm?id=2500365}, 140 | isbn = {978-1-4503-2326-0}, 141 | timestamp = {Thu, 19 Sep 2013 11:20:00 +0200}, 142 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/icfp/2013}, 143 | bibsource = {dblp computer science bibliography, http://dblp.org} 144 | } 145 | 146 | 147 | @inproceedings{DelahayeField, 148 | author = {David Delahaye and 149 | Micaela Mayero}, 150 | title = {{Field, une proc{\'{e}}dure de d{\'{e}}cision pour les nombres 151 | r{\'{e}}els en Coq}}, 152 | booktitle = {Journ{\'{e}}es francophones des langages applicatifs (JFLA'01)}, 153 | pages = {33--48}, 154 | year = {2001}, 155 | crossref = {DBLP:conf/jfla/2001}, 156 | timestamp = {Thu, 09 Feb 2006 12:05:13 +0100}, 157 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/jfla/DelahayeM01}, 158 | bibsource = {dblp computer science bibliography, http://dblp.org} 159 | } 160 | @proceedings{DBLP:conf/jfla/2001, 161 | editor = {Pierre Cast{\'{e}}ran}, 162 | title = {Journ{\'{e}}es francophones des langages applicatifs (JFLA'01), 163 | Pontarlier, France, Janvier, 2001}, 164 | series = {Collection Didactique}, 165 | publisher = {{INRIA}}, 166 | year = {2001}, 167 | isbn = {2-7261-1154-8}, 168 | timestamp = {Thu, 09 Feb 2006 12:05:13 +0100}, 169 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/jfla/2001}, 170 | bibsource = {dblp computer science bibliography, http://dblp.org} 171 | } 172 | 173 | 174 | @misc{Cregut04, 175 | TITLE = {{Une proc{\'{e}}dure de d{\'{e}}cision reflexive pour un fragment de l'arithm{\'{e}}tique de Presburger. In Journ{\'{e}}es Francophones des Langages Applicatifs.}}, 176 | AUTHOR = {P Cr{\'{e}}gut}, 177 | YEAR = {2004}, 178 | } 179 | 180 | 181 | @inproceedings{Lindblad04, 182 | author = {Lindblad, Fredrik and Benke, Marcin}, 183 | title = {A Tool for Automated Theorem Proving in Agda}, 184 | booktitle = {Proceedings of the 2004 International Conference on Types for Proofs and Programs}, 185 | series = {TYPES'04}, 186 | year = {2006}, 187 | isbn = {3-540-31428-8, 978-3-540-31428-8}, 188 | location = {Jouy-en-Josas, France}, 189 | pages = {154--169}, 190 | numpages = {16}, 191 | doi = {10.1007/11617990_10}, 192 | acmid = {2150519}, 193 | publisher = {Springer-Verlag}, 194 | address = {Berlin, Heidelberg}, 195 | } 196 | 197 | @inproceedings{Malecha14, 198 | author = {Gregory Malecha and 199 | Adam Chlipala and 200 | Thomas Braibant}, 201 | title = {Compositional Computational Reflection}, 202 | booktitle = {Interactive Theorem Proving - 5th International Conference, {ITP} 203 | 2014}, 204 | pages = {374--389}, 205 | year = {2014}, 206 | doi = {10.1007/978-3-319-08970-6_24}, 207 | timestamp = {Mon, 30 Jun 2014 12:30:02 +0200}, 208 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/itp/MalechaCB14}, 209 | bibsource = {dblp computer science bibliography, http://dblp.org} 210 | } 211 | @proceedings{DBLP:conf/itp/2014, 212 | editor = {Gerwin Klein and 213 | Ruben Gamboa}, 214 | title = {Interactive Theorem Proving - 5th International Conference, {ITP} 215 | 2014, Held as Part of the Vienna Summer of Logic, {VSL} 2014, Vienna, 216 | Austria, July 14-17, 2014. Proceedings}, 217 | series = {Lecture Notes in Computer Science}, 218 | volume = {8558}, 219 | publisher = {Springer}, 220 | year = {2014}, 221 | url = {http://dx.doi.org/10.1007/978-3-319-08970-6}, 222 | doi = {10.1007/978-3-319-08970-6}, 223 | isbn = {978-3-319-08969-0}, 224 | timestamp = {Mon, 30 Jun 2014 12:25:01 +0200}, 225 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/itp/2014}, 226 | bibsource = {dblp computer science bibliography, http://dblp.org} 227 | } 228 | 229 | @book{ChlipalaBook, 230 | author = {Adam Chlipala}, 231 | title = {Certified Programming with Dependent Types - {A} Pragmatic Introduction 232 | to the Coq Proof Assistant}, 233 | publisher = {{MIT} Press}, 234 | year = {2013}, 235 | isbn = {978-0-262-02665-9}, 236 | timestamp = {Mon, 11 May 2015 16:37:23 +0200}, 237 | biburl = {http://dblp.uni-trier.de/rec/bib/books/daglib/0035083}, 238 | bibsource = {dblp computer science bibliography, http://dblp.org} 239 | } 240 | 241 | 242 | @techreport{GonthierTuto, 243 | TITLE = {{An Ssreflect Tutorial}}, 244 | AUTHOR = {Gonthier, Georges and St{\'e}phane Le, Roux}, 245 | URL = {https://hal.inria.fr/inria-00407778}, 246 | TYPE = {Technical Report}, 247 | NUMBER = {RT-0367}, 248 | PAGES = {33}, 249 | INSTITUTION = {{INRIA}}, 250 | YEAR = {2009}, 251 | KEYWORDS = {proof assistants ; formal proofs ; Coq ; small scale reflection ; tactics}, 252 | PDF = {https://hal.inria.fr/inria-00407778/file/RT-367.pdf}, 253 | HAL_ID = {inria-00407778}, 254 | HAL_VERSION = {v1}, 255 | } 256 | 257 | @article{DelahayeLTac, 258 | author = {David Delahaye}, 259 | title = {A Proof Dedicated Meta-Language}, 260 | journal = {Electr. Notes Theor. Comput. Sci.}, 261 | volume = {70}, 262 | number = {2}, 263 | pages = {96--109}, 264 | year = {2002}, 265 | doi = {10.1016/S1571-0661(04)80508-5}, 266 | timestamp = {Tue, 21 Jun 2011 16:43:28 +0200}, 267 | biburl = {http://dblp.uni-trier.de/rec/bib/journals/entcs/Delahaye02}, 268 | bibsource = {dblp computer science bibliography, http://dblp.org} 269 | } 270 | 271 | @inproceedings{DelahayeModulo, 272 | author = {David Delahaye and 273 | Damien Doligez and 274 | Fr{\'{e}}d{\'{e}}ric Gilbert and 275 | Pierre Halmagrand and 276 | Olivier Hermant}, 277 | title = {Zenon Modulo: When Achilles Outruns the Tortoise Using Deduction Modulo}, 278 | booktitle = {Logic for Programming, Artificial Intelligence, and Reasoning - 19th 279 | International Conference, LPAR-19}, 280 | pages = {274--290}, 281 | year = {2013}, 282 | doi = {10.1007/978-3-642-45221-5_20}, 283 | timestamp = {Fri, 20 Dec 2013 14:41:17 +0100}, 284 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/lpar/DelahayeDGHH13}, 285 | bibsource = {dblp computer science bibliography, http://dblp.org} 286 | } 287 | @proceedings{DBLP:conf/lpar/2013, 288 | editor = {Kenneth L. McMillan and 289 | Aart Middeldorp and 290 | Andrei Voronkov}, 291 | title = {Logic for Programming, Artificial Intelligence, and Reasoning - 19th 292 | International Conference, LPAR-19, Stellenbosch, South Africa, December 293 | 14-19, 2013. Proceedings}, 294 | series = {Lecture Notes in Computer Science}, 295 | volume = {8312}, 296 | publisher = {Springer}, 297 | year = {2013}, 298 | url = {http://dx.doi.org/10.1007/978-3-642-45221-5}, 299 | doi = {10.1007/978-3-642-45221-5}, 300 | isbn = {978-3-642-45220-8}, 301 | timestamp = {Fri, 20 Dec 2013 14:36:05 +0100}, 302 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/lpar/2013}, 303 | bibsource = {dblp computer science bibliography, http://dblp.org} 304 | } 305 | 306 | @article{Leroy09, 307 | author = {Xavier Leroy}, 308 | title = {A formally verified compiler back-end}, 309 | journal = {CoRR}, 310 | volume = {abs/0902.2137}, 311 | year = {2009}, 312 | url = {http://arxiv.org/abs/0902.2137}, 313 | timestamp = {Mon, 05 Dec 2011 18:05:24 +0100}, 314 | biburl = {http://dblp.uni-trier.de/rec/bib/journals/corr/abs-0902-2137}, 315 | bibsource = {dblp computer science bibliography, http://dblp.org} 316 | } 317 | 318 | 319 | 320 | @inproceedings{DBLP:conf/mpc/KokkeS15, 321 | author = {Pepijn Kokke and 322 | Wouter Swierstra}, 323 | title = {Auto in {Agda} --- Programming Proof Search Using Reflection}, 324 | booktitle = {Mathematics of Program Construction - 12th International Conference, 325 | {MPC} 2015}, 326 | pages = {276--301}, 327 | year = {2015}, 328 | doi = {10.1007/978-3-319-19797-5_14}, 329 | timestamp = {Tue, 09 Jun 2015 13:13:03 +0200}, 330 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/mpc/KokkeS15}, 331 | bibsource = {dblp computer science bibliography, http://dblp.org} 332 | } 333 | 334 | @proceedings{DBLP:conf/mpc/2015, 335 | editor = {Ralf Hinze and 336 | Janis Voigtl{\"{a}}nder}, 337 | title = {Mathematics of Program Construction - 12th International Conference, 338 | {MPC} 2015, K{\"{o}}nigswinter, Germany, June 29 - July 1, 2015. 339 | Proceedings}, 340 | series = {Lecture Notes in Computer Science}, 341 | volume = {9129}, 342 | publisher = {Springer}, 343 | year = {2015}, 344 | url = {http://dx.doi.org/10.1007/978-3-319-19797-5}, 345 | doi = {10.1007/978-3-319-19797-5}, 346 | isbn = {978-3-319-19796-8}, 347 | timestamp = {Tue, 09 Jun 2015 13:11:48 +0200}, 348 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/mpc/2015}, 349 | bibsource = {dblp computer science bibliography, http://dblp.org} 350 | } 351 | 352 | 353 | @phdthesis{norell2007thesis, 354 | author = {Norell, Ulf}, 355 | file = {:Users/edwin/Documents/Mendeley Desktop/Norell - 2007 - Towards a practical programming language based on dependent type theory.pdf:pdf}, 356 | keywords = {Dependent Types}, 357 | mendeley-tags = {Dependent Types}, 358 | publisher = {Citeseer}, 359 | school = {Chalmers University of Technology}, 360 | title = {{Towards a practical programming language based on dependent type theory}}, 361 | year = {2007} 362 | } 363 | 364 | 365 | @inproceedings{DBLP:conf/tldi/Norell09, 366 | author = {Ulf Norell}, 367 | title = {Dependently typed programming in Agda}, 368 | booktitle = {Proceedings of TLDI'09: 2009 {ACM} {SIGPLAN} International Workshop 369 | on Types in Languages Design and Implementation}, 370 | pages = {1--2}, 371 | year = {2009}, 372 | doi = {10.1145/1481861.1481862}, 373 | timestamp = {Tue, 22 May 2012 15:24:55 +0200}, 374 | biburl = {http://dblp2.uni-trier.de/rec/bib/conf/tldi/Norell09}, 375 | bibsource = {dblp computer science bibliography, http://dblp.org} 376 | } 377 | 378 | @proceedings{DBLP:conf/tldi/2009, 379 | editor = {Andrew Kennedy and 380 | Amal Ahmed}, 381 | title = {Proceedings of TLDI'09: 2009 {ACM} {SIGPLAN} International Workshop 382 | on Types in Languages Design and Implementation, Savannah, GA, USA, 383 | January 24, 2009}, 384 | publisher = {{ACM}}, 385 | year = {2009}, 386 | url = {http://dl.acm.org/citation.cfm?id=1481861}, 387 | isbn = {978-1-60558-420-1}, 388 | timestamp = {Tue, 22 May 2012 15:24:55 +0200}, 389 | biburl = {http://dblp2.uni-trier.de/rec/bib/conf/tldi/2009}, 390 | bibsource = {dblp computer science bibliography, http://dblp.org} 391 | } 392 | 393 | -------------------------------------------------------------------------------- /paper-CICM-17/llncsdoc.sty: -------------------------------------------------------------------------------- 1 | % This is LLNCSDOC.STY the modification of the 2 | % LLNCS class file for the documentation of 3 | % the class itself. 4 | % 5 | \def\AmS{{\protect\usefont{OMS}{cmsy}{m}{n}% 6 | A\kern-.1667em\lower.5ex\hbox{M}\kern-.125emS}} 7 | \def\AmSTeX{{\protect\AmS-\protect\TeX}} 8 | % 9 | \def\ps@myheadings{\let\@mkboth\@gobbletwo 10 | \def\@oddhead{\hbox{}\hfil\small\rm\rightmark 11 | \qquad\thepage}% 12 | \def\@oddfoot{}\def\@evenhead{\small\rm\thepage\qquad 13 | \leftmark\hfil}% 14 | \def\@evenfoot{}\def\sectionmark##1{}\def\subsectionmark##1{}} 15 | \ps@myheadings 16 | % 17 | \setcounter{tocdepth}{2} 18 | % 19 | \renewcommand{\labelitemi}{--} 20 | \newenvironment{alpherate}% 21 | {\renewcommand{\labelenumi}{\alph{enumi})}\begin{enumerate}}% 22 | {\end{enumerate}\renewcommand{\labelenumi}{enumi}} 23 | % 24 | \def\bibauthoryear{\begingroup 25 | \def\thebibliography##1{\section*{References}% 26 | \small\list{}{\settowidth\labelwidth{}\leftmargin\parindent 27 | \itemindent=-\parindent 28 | \labelsep=\z@ 29 | \usecounter{enumi}}% 30 | \def\newblock{\hskip .11em plus .33em minus -.07em}% 31 | \sloppy 32 | \sfcode`\.=1000\relax}% 33 | \def\@cite##1{##1}% 34 | \def\@lbibitem[##1]##2{\item[]\if@filesw 35 | {\def\protect####1{\string ####1\space}\immediate 36 | \write\@auxout{\string\bibcite{##2}{##1}}}\fi\ignorespaces}% 37 | \begin{thebibliography}{} 38 | \bibitem[1982]{clar:eke3} Clarke, F., Ekeland, I.: Nonlinear 39 | oscillations and boundary-value problems for Hamiltonian systems. 40 | Arch. Rat. Mech. Anal. 78, 315--333 (1982) 41 | \end{thebibliography} 42 | \endgroup} 43 | -------------------------------------------------------------------------------- /paper-CICM-17/paper.tex: -------------------------------------------------------------------------------- 1 | % Franck Slama's and Edwin Brady's paper for CICM-2017 2 | % That follows the LLNCS style from Springer-Verlag 3 | % for Lecture Notes in Computer Science, 4 | % version 2.4 for LaTeX2e as of 16. April 2010 5 | % 6 | \documentclass{llncs} 7 | % 8 | \usepackage{makeidx} % allows for indexgeneration 9 | 10 | \usepackage{amsmath} 11 | \usepackage{amssymb} 12 | \usepackage{amsfonts} 13 | \usepackage{float} 14 | \usepackage{url} 15 | \usepackage{listings} 16 | 17 | \usepackage{tikz} 18 | \usetikzlibrary{matrix} 19 | 20 | \lstset{ 21 | basicstyle=\ttfamily, 22 | columns=fullflexible, 23 | keepspaces=true, 24 | mathescape 25 | } 26 | 27 | \def\code#1{{\mbox{\texttt{#1}}}} % For short element of code in the test that should not be hyphenated 28 | \def\codeL#1{{\texttt{#1}}} % For Long element of code in the text that might need to be hyphenated 29 | 30 | \newcommand{\figrule}{\begin{center}\hrule\end{center}} 31 | 32 | % 33 | \begin{document} 34 | % 35 | \pagestyle{headings} % switches on printing of running heads 36 | 37 | \mainmatter % start of the contributions 38 | % 39 | \title{Automatically Proving Equivalence by Type-Safe Reflection} 40 | %\subtitle{Or how to use dependent types to address some of the problems they have caused} 41 | % 42 | \titlerunning{xxx} % abbreviated title (for running head) 43 | % also used for the TOC unless 44 | % \toctitle is used 45 | % 46 | \author{Franck Slama \and Edwin Brady} 47 | \institute{University of St Andrews, Scotland, UK\\ 48 | \email{fs39@st-andrews.ac.uk, ecb10@st-andrews.ac.uk}} 49 | 50 | \maketitle % typeset the title of the contribution 51 | 52 | \begin{abstract} 53 | One difficulty with reasoning and programming with dependent types is that proof obligations arise naturally once programs become even moderately sized. For example, 54 | implementing an adder for binary numbers indexed over their natural number 55 | equivalents naturally leads to proof obligations for equalities of expressions 56 | over natural numbers. The need for these equality proofs comes, in intensional 57 | type theories, from the fact that the propositional equality enables us to 58 | prove as equal terms that are not judgementally equal, which means that the 59 | typechecker can't always obtain equalities by reduction. As far as possible, 60 | we would like to solve such proof obligations automatically. In this paper, we 61 | show one way to automate these proofs by reflection in the dependently typed 62 | programming language Idris. We show how defining reflected terms indexed by 63 | the original Idris expression allows us to construct and manipulate proofs. We 64 | build a hierarchy of tactics for proving equivalences in semi-groups, monoids, 65 | commutative monoids, groups, commutative groups, semi-rings and rings. We also 66 | show how each tactic reuses those from simpler structures, thus avoiding 67 | duplication of code and proofs. 68 | 69 | 70 | \keywords{proof automation, equivalence, equality, proof by reflection, correct-by-construction software, type-driven development} 71 | \end{abstract} 72 | % 73 | 74 | \input{./content/1-introduction.tex} 75 | 76 | % The chap 2 on the smaller tactic for Nat has been removed for space constraint 77 | 78 | \input{./content/3-hierarchyOfTactics.tex} 79 | 80 | \input{./content/4-relatedWork.tex} 81 | 82 | \input{./content/5-conclusions.tex} 83 | 84 | \subsubsection*{Acknowledgements} 85 | We thank the anonymous reviewers and Jacques Carette for their insightful 86 | comments on an earlier draft. We are also grateful for the support of 87 | the Scottish Informatics and Computer Science Alliance (SICSA) and 88 | EPSRC grant EP/N024222/1. 89 | 90 | \bibliographystyle{splncs03} 91 | \bibliography{dtp} 92 | 93 | \end{document} 94 | -------------------------------------------------------------------------------- /paper-CICM-17/rebuttal-proposed.txt: -------------------------------------------------------------------------------- 1 | We thank the reviewers for their careful reading of the paper and many helpful 2 | comments, which we will incorporate in the final draft. In particular, we are 3 | sure we can condense the paper according to Jacques' suggestions. 4 | 5 | Regarding the choice of structures and normal forms: our hierarchy is not 6 | supposed to deal with more than the most useful algebraic structures which 7 | arise in Idris programming. Coq for example only has a Ring solver but we have 8 | also provers for semigroups, monoids, and groups thanks to the modularity with 9 | which we compute the normal form. We focus on "pure" algebraic structures that 10 | do not have additional axioms and which do have a normal form. 11 | 12 | We don't claim that every interesting structure can be treated by one prover of 13 | our hierarchy of tactics, although our hierarchy of provers is more general than 14 | the automations available in e.g. Coq or Agda. 15 | 16 | On proving properties of normalisation: Things become difficult when we leave 17 | the ground of the "pure" algebraic structures in which we work (semigroups, monoids, 18 | groups and rings essentially) and move to an arbitrary equivalence relation with 19 | arbitrary axioms. Such a system can be seen as a rewriting system of deduction 20 | modulo. Determining if such a rewriting system admits a normal form requires 21 | determining whether the rewriting system is confluent and terminating. Both 22 | questions are difficult to answer in the general case. Currently, systems like 23 | Dedukti (an implementation of lambda-Pi-calculus modulo) can plug to an auxiliary 24 | confluence checker such as CSI^ho or ACPH. That would be an interesting future 25 | work for Idris, as suggested by the first reviewer. 26 | 27 | There are indeed many structures which do not have normal norms. But this problem, 28 | although very interesting, was somehow out of scope for our goals. We initially 29 | wanted Idris to be equipped with some automations, which we did in a modular way 30 | and with a new type-safe reflection approach, which made it easy to implement 31 | provers for other structures as well. 32 | 33 | It is interesting to note that the study of the properties of a proposed 34 | normalisation function can only be meta-theoretic. Indeed, Grégoire and 35 | Mahboubi noticed it for their Ring solver in Coq. However, this does not affect 36 | the correctness of the decision procedure. 37 | 38 | The motivating example is an illustration of how these proofs obligations occur 39 | naturally, but there are many other uses. 40 | -------------------------------------------------------------------------------- /paper-CICM-17/rebuttal_reviews_Brady_Slama_CICM17.txt: -------------------------------------------------------------------------------- 1 | We thank the reviews for their careful reading of the paper and many helpful 2 | comments, which we will address in a revision. In particular, we are sure we 3 | can condense the paper according to Jacques' useful suggestions. 4 | 5 | > although proofs of type equivalences are not the only use case 6 | 7 | Indeed; the motivating example is an illustration of how these proofs 8 | obligations occur naturally, but there are many other uses. 9 | 10 | > One question that comes to mind is whether it easy to prove properties of the 11 | > normalization functions 12 | 13 | Things become difficult when we leave the ground of the algebraic 14 | structures in which we work (monoids, groups and rings essentially) and 15 | move to an arbitrary equivalence relation with arbitrary axioms. Such a system 16 | can be seen as a rewriting system of deduction modulo. Determining if such a 17 | rewriting system admits a normal form requires determining whether the rewriting 18 | system is confluent and terminating. Both questions are difficult 19 | to answer in the general case. Currently, systems like Dedukti (an 20 | implementation of lambda-Pi-calculus modulo) can plug to an auxiliary 21 | confluence checker such as CSI^ho 22 | (http://cl-informatik.uibk.ac.at/software/csi/ho/) or ACPH 23 | (http://coco.nue.riec.tohoku.ac.jp/2015/papers/acph.pdf). 24 | 25 | It is interesting to note that the study of the properties of a proposed 26 | normalisation function can only be meta-theoretic: it's not possible to prove 27 | completeness within the language. Benjamin Grégoire and Assia Mahboubi 28 | [cited in our paper] have also noticed it for their implementation of their 29 | Ring prover for Coq. However, this does not affect the correctness of the 30 | decision procedure. 31 | 32 | > An interesting question for future work would be to exploit external provers 33 | 34 | Agreed, e.g. this would allow dealing with arbitrary rewriting rules. 35 | 36 | > ... as soon as the hierarchy grows large enough, such 37 | > 'natural' inheritance becomes very brittle .... 38 | 39 | Agreed, but our hierarchy is not supposed to deal with more than the most 40 | useful algebraic structures which arise in Idris programming. Coq for example 41 | only has a Ring solver but we have also provers for semigroups, monoids, and 42 | groups thanks to the modularity with which we compute the normal form. 43 | 44 | > there seems to be an implicit assumption that algebraic theories 45 | > have a canonical presentation 46 | 47 | To clarify: when we say groups admit a normal form, we mean structures that 48 | *only* have the axioms of a group. We focus on these "pure" algebraic 49 | structures (semi-groups, monoids, groups, rings...) that do not have 50 | additional axioms and which do have a normal form. 51 | 52 | > The largest flaw though seems to be the implicit assumption interesting structures have normal forms. 53 | 54 | Agreed, there are indeed many structures which do not have normal norms. But deduction modulo, 55 | although very interesting, was somehow out of scope for our goals. We initially wanted Idris 56 | to be equiped with some automations, which we do in a modular way which makes it easy to implement 57 | provers for other structures. 58 | 59 | We don't claim that every interesting structure can be treated by one prover of 60 | our hierachy of tactics, although our hierarchy of provers is more general than 61 | the automations available in e.g. Coq or Agda. 62 | 63 | > If one is going 64 | > to talk about a "correct by construction" approach, then the *definition* of what a 65 | > normal form is (i.e. syntax where intensional equality is a decision procedure for 66 | > equality). This is implicit everywhere, but never explicit. And it is not an 67 | > invariant kept by the types! 68 | 69 | Yes, the types only say that the original reflected term and the normalised one both represent the 70 | same concrete value, and this is in fact the proof of *correctness* of the approach. 71 | The *completeness* is a meta property that we cannot prove in the language itself. 72 | Benjamin Grégoire and Assia Mahboubi [cited in our paper] have also noticed it for their implementation 73 | of their Ring prover for Coq. It's not an issue as it does not affect the correctness of the decision procedure. 74 | 75 | > For example, the word problem for groups is undecidable. This implies that there are 76 | > theories with no normal forms. 77 | 78 | Yes, we agree, see our answer to the two previous questions. 79 | 80 | > In section 2.5, an *order* is snuck in, onto the set of variables. This is crucial! 81 | > It should not be buried. Hint: Groebner bases. Term orders are important. 82 | > But normal forms are extremely sensitive to term orders (as bad as exponential). 83 | 84 | The set of variables V is said to be ordered in section 2.5 (page 12). Note that this is not a limitation 85 | as it is always possible to build an adhoc ordering by using the order of appearance of the variables in the goal A = B 86 | This is in fact exactly what our automatically reflection does (mentionned briefly in section 2.7). 87 | 88 | > Furthermore, the discussion in this section implicitly assumes associativity. There 89 | > are many useful, interesting, non-associative algebras... 90 | 91 | We agree that there are many useful non-associative algebras. However, our work 92 | was focused on rings and all of the underneath structures. This is why, at the 93 | very root, we have semi-groups. Therefore, the associativity is the first axiom 94 | that we deal with, and our machinery is not intended for non-associative 95 | algebras. 96 | 97 | > buried assumption, that all terms over algebraic structures have decidable equality! 98 | 99 | To clarify: the required decidable equality is only a decidable equality for 100 | the constants, e.g, when proving equalities over Nat, the needed decidable 101 | equality is the one that decides the equality between two natural numbers (with 102 | no variables and no operations), called decEq in Idris. 103 | 104 | > - The interface to a Group should have *either* unary or binary -, and then use a 105 | > *definitional* (and thus conservative) extensio to add the other. This is again 106 | > what Realms are useful for. 107 | 108 | Having both is useful if the user has something to prove with both the unary Neg and the binary Minus. 109 | Expressing in the interface something like "there is either a unary or binary -" 110 | would have been a hassle. However, this is not a limitation : it is always possible 111 | for the user to define one with the other (and the proof Minus_simpl that he/she 112 | has to provide then becomes trivial). 113 | 114 | > this isn't something you have a choice to do / not do. 115 | 116 | Indeed, we *must* do it. We still wanted to describe what the normalisation function has to do, 117 | and not to only mention the arbitrary decisions that we had to take (such as writing everything on 118 | the complete left or complete right associative form). 119 | 120 | These simplifications entail simplifying the addition with zero, and the multiplication 121 | with the constants zero and with one, doing the computations between two nearby constants, etc. 122 | (mentionned briefly on page 13) 123 | 124 | > The hacks involved in Variable (i.e. treat Neg as a pseudo-variable) don't scale 125 | > well. 126 | 127 | We agree, but our goal was to implement automatic provers for the albebraic 128 | structures that are underneath rings, not to deal with all possible algebraic 129 | structure. We'll make this clearer in the paper. 130 | 131 | 132 | -------------------------------------------------------------------------------- /paper-CICM-17/remreset.sty: -------------------------------------------------------------------------------- 1 | 2 | % remreset package 3 | %%%%%%%%%%%%%%%%%% 4 | 5 | % Copyright 1997 David carlisle 6 | % This file may be distributed under the terms of the LPPL. 7 | % See 00readme.txt for details. 8 | 9 | % 1997/09/28 David Carlisle 10 | 11 | % LaTeX includes a command \@addtoreset that is used to declare that 12 | % a counter should be reset every time a second counter is incremented. 13 | 14 | % For example the book class has a line 15 | % \@addtoreset{footnote}{chapter} 16 | % So that the footnote counter is reset each chapter. 17 | 18 | % If you wish to bas a new class on book, but without this counter 19 | % being reset, then standard LaTeX gives no simple mechanism to do 20 | % this. 21 | 22 | % This package defines |\@removefromreset| which just undoes the effect 23 | % of \@addtorest. So for example a class file may be defined by 24 | 25 | % \LoadClass{book} 26 | % \@removefromreset{footnote}{chapter} 27 | 28 | 29 | \def\@removefromreset#1#2{{% 30 | \expandafter\let\csname c@#1\endcsname\@removefromreset 31 | \def\@elt##1{% 32 | \expandafter\ifx\csname c@##1\endcsname\@removefromreset 33 | \else 34 | \noexpand\@elt{##1}% 35 | \fi}% 36 | \expandafter\xdef\csname cl@#2\endcsname{% 37 | \csname cl@#2\endcsname}}} 38 | 39 | 40 | -------------------------------------------------------------------------------- /paper-CICM-17/sprmindx.sty: -------------------------------------------------------------------------------- 1 | delim_0 "\\idxquad " 2 | delim_1 "\\idxquad " 3 | delim_2 "\\idxquad " 4 | delim_n ",\\," 5 | --------------------------------------------------------------------------------