├── .gitignore ├── EggTactic.lean ├── EggTactic ├── Egraph.lean ├── Sexp.lean ├── Test.lean └── Tracing.lean ├── Evaluation ├── Calc.lean ├── FunctionalProgramming.lean ├── GeneralizedRewriting.lean ├── GroupsKnuthBendix.lean ├── GroupsKnuthBendix.v ├── GroupsKnuthBendixAesop.lean ├── GroupsKnuthBendixAuto.v ├── GroupsKnuthBendixExploded.v ├── GroupsKnuthBendixSimp.lean ├── ICJAR.lean ├── Reviewer1 │ ├── counts.sh │ └── gen_count.py ├── scaling-proof-length │ ├── .gitignore │ ├── Rplots.pdf │ ├── Scaling.lean │ ├── Scaling.v │ ├── plotscaling.R │ ├── scaling-exponential.csv │ ├── scaling-exponential.png │ ├── scaling-exponential.py │ ├── scaling-exponential.tex │ ├── scaling-geometric.csv │ ├── scaling-geometric.py │ └── stats.csv ├── scaling-search-space │ ├── plotscaling.R │ └── scaling-geometric.py └── scaling-space-and-proof │ ├── .gitignore │ ├── Rplots.pdf │ ├── Scaling.lean │ ├── Scaling.v │ ├── plotscaling.R │ ├── scaling-exponential.csv │ ├── scaling-exponential.png │ ├── scaling-exponential.py │ ├── scaling-exponential.tex │ ├── scaling-geometric.csv │ ├── scaling-geometric.png │ ├── scaling-geometric.py │ ├── scaling-geometric.tex │ └── stats.csv ├── ForeignLean ├── .gitignore ├── Main.lean ├── Makefile └── myfuns.cpp ├── Main.lean ├── PrettyPrinter.lean ├── README.md ├── example.lean ├── ffi-egg ├── .gitignore ├── Cargo.toml └── src │ ├── lib.rs │ ├── main.rs │ ├── math.rs │ └── rules.rs ├── json-egg ├── .gitignore ├── Cargo.lock ├── Cargo.toml ├── compare-json-reqs.py ├── json-inputs │ ├── egg.json │ ├── group-inv-egg-cleaned-up.json │ ├── group-inv-egg-not-cleaned-up.json │ ├── group-inv-inv-broken-rewrites.json │ ├── group-inv-inv-working-rewrites.json │ ├── group-inv-mul-cancel-left.json │ ├── group-inv-mul-slow-reduced.json │ ├── group-inv-mul-slow.json │ ├── group-inv-mul-working.json │ ├── group-inv-simplified.json │ ├── group-test-ap.json │ ├── group-test.json │ ├── instantiation-test.json │ ├── non-equal-test.json │ ├── slow.json │ ├── test-position-vars.json │ ├── test-position.json │ ├── test.json │ ├── transitive-rev-test.json │ └── transitive-test.json ├── solution.md └── src │ ├── main.rs │ └── scheduler.rs ├── lake-manifest.json ├── lakefile.lean └── lean-toolchain /.gitignore: -------------------------------------------------------------------------------- 1 | *.code-* 2 | /target 3 | /build 4 | /lean_packages 5 | tags 6 | *.code-workspace 7 | lake-packages/ 8 | lean_packages/ 9 | *~ 10 | \#* 11 | .#* 12 | *.lock 13 | build 14 | GPATH 15 | GRTAGS 16 | GSYMS 17 | GTAGS 18 | .projectile 19 | .lean_options 20 | .vs 21 | compile_commands.json 22 | *.idea 23 | tasks.json 24 | settings.json 25 | .gdb_history 26 | .vscode 27 | *.produced.out 28 | CMakeSettings.json 29 | CppProperties.json 30 | tags 31 | utils/ 32 | -------------------------------------------------------------------------------- /EggTactic/Egraph.lean: -------------------------------------------------------------------------------- 1 | import Std.Data.HashMap 2 | open Std HashMap 3 | 4 | namespace Egraph 5 | 6 | -- We key pointers with names instead of types so that 7 | -- we can have names of structures that does not create a cycle. 8 | -- When defining a structure, we are not allowed to use the 9 | -- type of the structure in the body of the structure. 10 | -- We dodge this restriction by converting to names. 11 | -- Convention: all pointer variables will have name starting 12 | -- with π. e.g. πx is a pointer to x. 13 | structure Ptr (_α : Name) where ix : UInt64 14 | deriving DecidableEq, Hashable,Inhabited 15 | 16 | -- typeclass that can dereference points of type 'o' named by 'α' 17 | class DerefM (m : Type → Type) (α : Name) (o : outParam Type) where 18 | deref : Ptr α -> m (Option o) 19 | set : Ptr α -> o -> m Unit 20 | malloc [Inhabited o] : m (Ptr α) -- allocate junk memory 21 | 22 | def DerefM.new [Monad m] [DerefM m α o] [Inhabited o] (v : o) : m (Ptr α) := do 23 | let ptr ← malloc 24 | set ptr v 25 | return ptr 26 | 27 | -- dereference pointer and run action. action must succeed. 28 | def Ptr.modifyM! [Monad m] [DerefM m α o] [Inhabited o] 29 | (f : o → m o) (p : Ptr α) : m Unit := do 30 | let v ← Option.get! <$> (DerefM.deref p) 31 | DerefM.set p (← f v) 32 | 33 | def Ptr.modify! [Monad m] [DerefM m α o] [Inhabited o] 34 | (f : o → o) (p : Ptr α) : m Unit := 35 | p.modifyM! (fun o => pure (f o)) 36 | 37 | def Ptr.deref [DerefM m α o] (p : Ptr α) : m (Option o) := DerefM.deref p 38 | def Ptr.deref! [Inhabited o] [Monad m] [DerefM m α o] (p : Ptr α) : m o:= 39 | Option.get! <$> DerefM.deref p 40 | 41 | abbrev Memory (α : Type) := HashMap UInt64 α 42 | notation "*?" x => DerefM.deref x -- dereference a pointer 43 | set_option quotPrecheck false in 44 | notation "*" x => (← Ptr.deref! x) -- dereference a pointer 45 | notation p " ^= " v => DerefM.set p v -- set a pointer to a value 46 | -- notation "^" x => Ptr `x 47 | 48 | -- terms. 49 | structure HashConsTerm (σ : Type) where 50 | head : σ 51 | args : Array (Ptr `Egraph.Eclass) 52 | deriving BEq, DecidableEq, Hashable, Inhabited 53 | 54 | -- equivalence class of terms. 55 | structure Eclass (σ : Type) where 56 | members : Array (HashConsTerm σ) 57 | -- TODO: cannot use (Ptr Eclass) here :( 58 | -- lmao, just use names! 59 | parentTerms : Array (HashConsTerm σ × Ptr `Egraph.Eclass) 60 | πcanon : Ptr `Egraph.Eclass -- pointer to canonical eclass represenative 61 | subtreeSize : UInt64 -- union find subtree size 62 | deriving BEq, DecidableEq, Hashable, Inhabited 63 | 64 | -- create a singleton e class. 65 | def Eclass.singleton (πcanon : Ptr ``Eclass) (member : HashConsTerm σ) : Eclass σ where 66 | members := #[member] 67 | parentTerms := {} 68 | πcanon := πcanon 69 | subtreeSize := 0 70 | 71 | -- add a term and the e-class that owns the term as the parent 72 | -- of this e-class. 73 | -- we add the parent e-class as well as the term since this informatoin 74 | -- is useful when moving terms around during the update. 75 | def Eclass.addParent (cls : Eclass σ) 76 | (tm : HashConsTerm σ) 77 | (πtm_cls : Ptr ``Eclass) := 78 | { cls with parentTerms := cls.parentTerms.push (tm, πtm_cls) } 79 | 80 | -- the data of the Egraph state. 81 | structure State (σ : Type) where 82 | σinhabited : Inhabited σ 83 | σbeq : DecidableEq σ 84 | σhashable : Hashable σ 85 | eclasses : Memory (Eclass σ) 86 | hashconsterms : Memory (HashConsTerm σ) 87 | term2classp : HashMap (HashConsTerm σ) (Ptr ``Eclass) 88 | 89 | def State.new [Inhabited σ] [DecidableEq σ] [Hashable σ] : State σ where 90 | σinhabited := inferInstance 91 | σbeq := inferInstance 92 | σhashable := inferInstance 93 | eclasses := {} 94 | hashconsterms := {} 95 | term2classp := {} 96 | 97 | abbrev M σ α := StateM (State σ) α 98 | 99 | -- deref an e-class 100 | instance : DerefM (M σ) ``Eclass (Eclass σ) where 101 | deref ptr := do 102 | return (← get).eclasses.find? ptr.ix 103 | set ptr v := modify (fun s => 104 | { s with eclasses := s.eclasses.insert ptr.ix v : State σ }) 105 | malloc := do 106 | let ptr := Ptr.mk <| UInt64.ofNat (← get).eclasses.size 107 | return ptr 108 | 109 | instance : DerefM (M σ) ``HashConsTerm (HashConsTerm σ) where 110 | deref ptr := do 111 | return (← get).hashconsterms.find? ptr.ix 112 | set ptr v := modify (fun s => 113 | { s with hashconsterms := s.hashconsterms.insert ptr.ix v : State σ }) 114 | malloc := do 115 | let ptr := Ptr.mk <| UInt64.ofNat (← get).eclasses.size 116 | return ptr 117 | 118 | class Canonicalize (σ : Type) (α : Type) where 119 | canonicalize : α → M σ α 120 | 121 | -- return canonical pointer 122 | notation "⟦" x "⟧" => Canonicalize.canonicalize x 123 | 124 | partial def canonicalizeClass (πcls : Ptr ``Eclass) : M σ (Ptr ``Eclass) := do 125 | let cls : (Eclass σ) ← πcls.deref! 126 | if cls.πcanon == πcls 127 | then return πcls 128 | else do 129 | let πcanon ← Egraph.canonicalizeClass πcls 130 | πcls ^= { cls with πcanon := πcanon } 131 | return πcanon 132 | 133 | instance : Canonicalize σ (Ptr ``Eclass) where 134 | canonicalize := canonicalizeClass 135 | 136 | partial def canonicalizeHashConsTerm (htm : HashConsTerm σ) : 137 | M σ (HashConsTerm σ) := do 138 | let _ : Inhabited σ := (← get).σinhabited 139 | let htm := { htm with args := (← htm.args.mapM canonicalizeClass) } 140 | return htm 141 | 142 | instance : Canonicalize σ (HashConsTerm σ) where 143 | canonicalize := canonicalizeHashConsTerm 144 | 145 | partial def HashConsTerm.find! (htm : HashConsTerm σ) : 146 | M σ (Ptr ``Eclass) := do 147 | match (← get).term2classp.find? (← ⟦htm⟧) with 148 | | .some cls => return cls 149 | | .none => panic! "unable to find e-class" 150 | 151 | partial def HashConsTerm.findOrAdd (htm : HashConsTerm σ) : 152 | M σ (Ptr ``Eclass) := do 153 | let htm ← ⟦htm⟧ 154 | let πhtm_cls ← 155 | match (← get).term2classp.find? htm with 156 | | .some x => pure x 157 | | .none => do 158 | let πhtm_cls ← DerefM.malloc (α := ``Eclass) 159 | modify (fun state => { state with term2classp := state.term2classp.insert htm πhtm_cls }) 160 | πhtm_cls ^= Eclass.singleton (πcanon := πhtm_cls) (member := htm) 161 | return πhtm_cls 162 | -- for each argument, update the parent e-elcass by adding a pointer to the 163 | -- e class of htm 164 | for πarg in htm.args do 165 | πarg.modify! (fun cls => cls.addParent htm πhtm_cls) 166 | return πhtm_cls 167 | 168 | mutual 169 | 170 | -- rebuild E-graph 171 | partial def Egraph.unite (πc πd : Ptr ``Eclass) : M σ (Ptr ``Eclass) := do 172 | let πc ← ⟦ πc ⟧ 173 | let πd ← ⟦ πd ⟧ 174 | if πc == πd then return πc 175 | -- attach root of lighter thing to root of heavier thing, so that there is 176 | -- a chance that depth does not go up. 177 | let (πparent, πchild) := 178 | if (← πc.deref!).subtreeSize >= (← πd.deref!).subtreeSize 179 | then (πc, πd) 180 | else (πd, πc) 181 | 182 | πchild.modify! (fun c => { c with πcanon := πparent }) 183 | πparent.modifyM! (fun c => do return { 184 | c with subtreeSize := c.subtreeSize + (← πchild.deref!).subtreeSize 185 | }) 186 | πparent.modifyM! (fun c => do return { 187 | c with 188 | members := c.members ++ (← πchild.deref!).members, 189 | parentTerms := c.parentTerms ++ (← πchild.deref!).parentTerms 190 | }) 191 | πchild.modify! (fun c => { c with members := #[], parentTerms := #[] }) 192 | Egraph.rebuild πparent 193 | return πparent 194 | 195 | -- rebuild E-class. 196 | partial def Egraph.rebuild (πc : Ptr ``Eclass) : M σ Unit := do 197 | let mut parents := #[] 198 | for (htm, πhtm_old_cls) in (← πc.deref!).parentTerms do 199 | let πhtm_new_cls ← htm.findOrAdd 200 | let πhtm_new_cls ← Egraph.unite πhtm_old_cls πhtm_new_cls 201 | modify (fun s => { s with term2classp := (s.term2classp.erase htm)} ) 202 | let htm ← ⟦ htm ⟧ 203 | modify (fun s => { s with term2classp := s.term2classp.insert htm πhtm_new_cls } ) 204 | parents := parents.push (htm, πhtm_new_cls) 205 | πc.modify! (fun c => { c with parentTerms := parents}) 206 | end 207 | -------------------------------------------------------------------------------- /EggTactic/Test.lean: -------------------------------------------------------------------------------- 1 | import EggTactic 2 | 3 | set_option trace.EggTactic.egg true 4 | -- TODO @andres: what egg gives us back is wrong. What we really need to know is the rewrite term that we need to use. 5 | -- This means we should build the sexpr ( ... ) and then use this to perform the rewrite. 6 | -- Instead, what `egg` gives us is the goal state after the rewrite. 7 | -- 8 | -- This is a problem because we need to know the rewrite term to perform the rewrite. 9 | -- We can't just use the goal state because the goal state is the result of the rewrite. 10 | 11 | 12 | theorem testSuccess0 (anat: Nat) (bnat: Nat) (H: anat = bnat): anat = bnat := by { 13 | intros; 14 | eggxplosion [H] 15 | } 16 | 17 | #print testSuccess0 18 | 19 | 20 | set_option pp.analyze true in 21 | theorem testSuccess (anat: Nat) (bint: Int) (cnat: Nat) 22 | (dint: Int) (eint: Int) (a_eq_a: anat = anat) (b_eq_d: bint = dint) (d_eq_e: dint = eint): bint = eint := by 23 | -- eggxplosion [not_rewrite] 24 | -- eggxplosion [rewrite_wrong_type] 25 | -- rewrite [b_eq_d] 26 | -- rewrite [d_eq_e] 27 | -- rfl 28 | eggxplosion [b_eq_d, d_eq_e] 29 | 30 | 31 | #print testSuccess 32 | 33 | 34 | theorem testSuccess0Symm (anat: Nat) (bnat: Nat) (H: bnat = anat): anat = bnat := by { 35 | intros; 36 | eggxplosion [H] 37 | } 38 | 39 | #print testSuccess0Symm 40 | 41 | -- elab "boiledEgg" "[" rewrites:ident,* "]" : tactic => withMainContext do 42 | 43 | -- | test that we can run rewrites 44 | 45 | 46 | 47 | -- | test that we can run theorems in reverse. 48 | theorem testSuccessRev : ∀ (anat: Nat) (bint: Int) (cnat: Nat) 49 | (dint: Int) (eint: Int) (a_eq_a: anat = anat) (b_eq_d: bint = dint) (d_eq_e: dint = eint), 50 | eint = bint := by 51 | intros a b c d e aeqa beqd deqe 52 | eggxplosion [beqd, deqe] 53 | 54 | #print testSuccessRev 55 | 56 | 57 | 58 | /- 59 | theorem testManualInstantiation 60 | (group_inv: forall (g: Int), g - g = 0) 61 | (h: Int) (k: Int): h - h = k - k := by 62 | have gh : h - h = 0 := group_inv h 63 | have gk : k - k = 0 := group_inv k 64 | -- TODO: figure out how to get these 65 | -- manually instantiated locals to work. 66 | -- @bollu's hypothesis is that we need to force 67 | -- metavariable resolution at the value and type level 68 | -- with a couple 'reduce's. 69 | eggxplosion [gh, gk] 70 | 71 | #print testManualInstantiation 72 | -/ 73 | 74 | 75 | theorem assoc_instantiate(G: Type) 76 | (mul: G → G → G) 77 | (assocMul: forall (a b c: G), (mul (mul a b) c) = mul a (mul b c)) 78 | (x y z: G) : mul x (mul y z) = mul (mul x y) z := by { 79 | eggxplosion [assocMul] 80 | } 81 | 82 | #print assoc_instantiate 83 | 84 | 85 | /- 86 | theorem testGoalNotEqualityMustFail : ∀ (a: Nat) (b: Int) (c: Nat) , Nat := by 87 | intros a b c 88 | eggxplosion [] 89 | sorry 90 | -/ 91 | 92 | 93 | 94 | /- 95 | rw!("assoc-mul"; "(* ?a (* ?b ?c))" => "(* (* ?a ?b) ?c)"), 96 | rw!("assoc-mul'"; "(* (* ?a ?b) ?c)" => "(* ?a (* ?b ?c))"), 97 | rw!("one-mul"; "(* 1 ?a)" => "?a"), 98 | rw!("one-mul'"; "?a" => "(* 1 ?a)"), 99 | rw!("inv-left"; "(* (^-1 ?a) ?a)" => "1"), 100 | rw!("inv-left'"; "1" => "(* (^-1 a) a)"), 101 | rw!("inv-left''"; "1" => "(* (^-1 b) b)"), 102 | rw!("mul-one"; "(* ?a 1)" => "?a"), 103 | rw!("mul-one'"; "?a" => "(* ?a 1)" ), 104 | rw!("inv-right"; "(* ?a (^-1 ?a))" => "1"), 105 | rw!("inv-right'"; "1" => "(* a (^-1 a))"), 106 | rw!("inv-right''"; "1" => "(* b (^-1 b))"), 107 | //rw!("anwser''"; "(* (^-1 b)(^-1 a))" => "ANSWER"), 108 | -/ 109 | theorem inv_mul_cancel_left (G: Type) 110 | (inv: G → G) 111 | (mul: G → G → G) 112 | (one: G) 113 | (x: G) 114 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 115 | (invLeft: forall (a: G), mul (inv a) a = one) 116 | (mulOne: forall (a: G), a = mul a one) 117 | (invRight: forall (a: G), one = mul a (inv a)) 118 | --(invRightX: one = mul x (inv x)) 119 | : (inv (inv x) = x) := by 120 | eggxplosion [assocMul, invLeft, mulOne, invRight] (timeLimit := 3) 121 | 122 | #print inv_mul_cancel_left 123 | 124 | def all {α : Type} (p : α → Bool) (xs : List α) := List.and (List.map p xs) 125 | def all'{α : Type} (p : α → Bool) (xs : List α) := List.foldr (λ a b => b && (p a)) True xs 126 | 127 | theorem deforestation : ∀ (p : α → Bool) (xs : List α), all p xs = all' p xs := by 128 | intros p xs 129 | unfold all 130 | unfold all' 131 | unfold List.map 132 | unfold List.and 133 | unfold List.all 134 | induction xs with 135 | | nil => 136 | simp 137 | rfl 138 | | cons head tail ih => 139 | -- ⊢ List.foldr (fun a r => a && r) true (p head :: List.map p tail) = List.foldr (fun a b => b && p a) true (head :: tail) 140 | 141 | 142 | 143 | 144 | theorem testInstantiation 145 | -- TODO: need to change definitions to make all arguments implicit, since those are the only kinds of rewrites 146 | -- egg can cope with! 147 | (group_inv: forall {g: Int}, g - g = 0) 148 | (h: Int) (k: Int): h - h = k - k := by { 149 | -- expected no bound variables, we use locally nameless. 150 | -- @andres: I strongly suspect the toExplode 151 | -- array somehow leaking in `bvars` on accident. A cursory glance at it did not show me what it 152 | -- does when it doesn't explode a variable; I would have expected it to instantiate an `mvar`. 153 | eggxplosion [group_inv]; 154 | } 155 | 156 | 157 | #print testInstantiation 158 | -------------------------------------------------------------------------------- /EggTactic/Tracing.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | initialize Lean.registerTraceClass `EggTactic.egg 3 | -------------------------------------------------------------------------------- /Evaluation/Calc.lean: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/opencompl/egg-tactic-code/8b4aa748047a43213fc2c0dfca6b7af4a475f785/Evaluation/Calc.lean -------------------------------------------------------------------------------- /Evaluation/FunctionalProgramming.lean: -------------------------------------------------------------------------------- 1 | import EggTactic 2 | namespace Test 3 | 4 | 5 | inductive List α 6 | | nil : List α 7 | | cons : α → List α → List α 8 | 9 | 10 | notation a "::" b => List.cons a b 11 | notation "[" "]" => List.nil 12 | 13 | -- From init/data/list/basic.lean 14 | 15 | def append : List α → List α → List α 16 | | [], bs => bs 17 | | a::as, bs => a :: append as bs 18 | 19 | /- I should be able to derive these directly in the tactic... -/ 20 | def append_nil : ∀ (bs: List α), append ([] : List α) bs = bs := by 21 | intros 22 | rfl 23 | 24 | def append_cons : ∀ (a : α) (as bs : List α), append (a :: as) bs = a :: (append as bs) := by 25 | intros 26 | rfl 27 | 28 | instance {α : Type _} : HAppend (List α) (List α) (List α) where hAppend := append 29 | 30 | theorem append_assoc (as bs cs : List α) : (as ++ bs) ++ cs = as ++ (bs ++ cs) := by 31 | induction as with 32 | | nil => eggxplosion [append_nil, append_cons] 33 | | cons a as ih => eggxplosion [ih, append_nil, append_cons] -- could this also be automated? 34 | -- <;> eggxplosion 35 | 36 | 37 | def reverseAux : List α → List α → List α 38 | | [], r => r 39 | | a::l, r => reverseAux l (a::r) 40 | 41 | def List.reverse (as : List α) :List α := 42 | reverseAux as [] 43 | 44 | def reverseAux_nil : ∀ r : List α, reverseAux ([] : List α) r = r := by 45 | intros 46 | rfl 47 | 48 | def reverseAux_cons : ∀ (l r : List α) (a : α) , reverseAux (a :: l) r = reverseAux l (a :: r) := by 49 | intros 50 | rfl 51 | 52 | def reverse_def : ∀ l : List α, l.reverse = reverseAux l [] := by 53 | intros 54 | rfl 55 | 56 | theorem reverseAux_eq_append (as bs : List α) : reverseAux as bs = reverseAux as [] ++ bs := by 57 | induction as generalizing bs with 58 | | nil => eggxplosion [reverseAux_nil, reverseAux_cons] 59 | | cons a as ih => 60 | eggxplosion [reverseAux_nil, reverseAux_cons, append_assoc] 61 | -- <;> eggxplosion [reverseAux, append_assoc] 62 | 63 | theorem reverse_cons (a : α) (as : List α) : List.reverse (a :: as) = List.reverse as ++ (a :: List.nil) := by 64 | eggxplosion [reverse_def, reverseAux_nil, reverseAux_cons, reverseAux_eq_append] 65 | 66 | theorem reverse_append (as bs : List α) : (as ++ bs).reverse = bs.reverse ++ as.reverse := by 67 | induction as generalizing bs with 68 | | nil => eggxplosion [] 69 | | cons a as ih => eggxplosion [ih, append_assoc] 70 | -- <;> eggxplosion [append_assoc] 71 | 72 | def List.map : (α → β) → List α → List β 73 | | _, [] => [] 74 | | f, (a :: as) => (f a) :: (as.map f) 75 | 76 | def List.foldr (f : α → β → β) (init : β) : List α → β 77 | | [] => init 78 | | a :: l => f a (foldr f init l) 79 | 80 | -- A short cut to deforestation, Gill, Launchbury, Peyton Jones 81 | def all (p : α → Bool) (xs : List α) : Bool := (List.map p xs).foldr and true 82 | 83 | def all' : (α → Bool) → List α → Bool 84 | | _, [] => true 85 | | p, (x::xs) => (p x) && (all' p xs) 86 | 87 | theorem all_deforestation {α} : ∀ (p : α → Bool) (xs : List α), all p xs = all' p xs := by 88 | -- eggxplosion [all, all', map, foldr] 89 | intros p xs 90 | try induction xs with 91 | | nil => simp [all, all', List.foldr, List.map] 92 | | cons a as ih => simp [all, all', List.foldr, List.map, ih] 93 | eggxplosion [all, all', List.foldr, List.map] -- add induction explicitly (like in Agda?) 94 | -------------------------------------------------------------------------------- /Evaluation/GeneralizedRewriting.lean: -------------------------------------------------------------------------------- 1 | import EggTactic 2 | 3 | inductive EQ {α : Type} (a : α) : α → Type where 4 | | refl : EQ a a 5 | 6 | def EQ.trans (h₁ : EQ a b) (h₂ : EQ b c) : EQ a c := by 7 | cases h₁; cases h₂; constructor 8 | 9 | instance : Trans (@EQ α) (@EQ α) (@EQ α) where 10 | trans := EQ.trans 11 | 12 | infix:50 " ≋ " => EQ 13 | 14 | example (h₁ : EQ a b) (h₂ : b = c) (h₃ : EQ c d) : EQ a d := by 15 | calc a ≋ b := h₁ 16 | _ = c := h₂ 17 | _ ≋ d := h₃ 18 | 19 | example (h₁ : EQ a b) (h₂ : b = c) (h₃ : EQ c d) : EQ a d := by 20 | eggxplosion [h₁, h₂, h₃] 21 | -------------------------------------------------------------------------------- /Evaluation/GroupsKnuthBendix.lean: -------------------------------------------------------------------------------- 1 | import EggTactic 2 | 3 | set_option trace.EggTactic.egg true 4 | theorem inv_inv 5 | {G: Type} 6 | (inv: G → G) 7 | (mul: G → G → G) 8 | (one: G) 9 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 10 | (invLeft: forall (a: G), mul (inv a) a = one) 11 | (oneMul: forall (a: G), mul one a = a) 12 | (mulOne: forall (a: G), mul a one = a) 13 | (invRight: forall (a: G), mul a (inv a) = one) 14 | (x: G) 15 | : (inv (inv x) = x) := by 16 | eggxplosion [assocMul, invLeft, mulOne, oneMul, invRight] 17 | 18 | theorem inv_mul_cancel_left 19 | {G: Type} 20 | (inv: G → G) 21 | (mul: G → G → G) 22 | (one: G) 23 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 24 | (invLeft: forall (a: G), mul (inv a) a = one) 25 | (mulOne: forall (a: G), mul a one = a) 26 | (oneMul: forall (a: G), mul one a = a) 27 | (invRight: forall (a: G), mul a (inv a) = one) 28 | (x y : G) 29 | : (mul (inv x) (mul x y)) = y := by 30 | eggxplosion [assocMul, invLeft, mulOne, oneMul, invRight] 31 | 32 | 33 | theorem mul_inv_cancel_left 34 | {G: Type} 35 | (inv: G → G) 36 | (mul: G → G → G) 37 | (one: G) 38 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 39 | (invLeft: forall (a: G), mul (inv a) a = one) 40 | (mulOne: forall (a: G), a = mul a one) 41 | (oneMul: forall (a: G), mul one a = a) 42 | (invRight: forall (a: G), one = mul a (inv a)) 43 | (x y : G) 44 | : (mul x (mul (inv x) y)) = y := by 45 | eggxplosion [assocMul, invLeft, mulOne, oneMul, invRight] 46 | 47 | 48 | theorem inv_mul 49 | {G: Type} 50 | (inv: G → G) 51 | (mul: G → G → G) 52 | (one: G) 53 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 54 | (invLeft: forall (a: G), mul (inv a) a = one) 55 | (oneMul: forall (a: G), mul one a = a) 56 | (mulOne: forall (a: G), mul a one = a) 57 | (invRight: forall (a: G), mul a (inv a) = one) 58 | (x y : G) 59 | : (inv (mul x y)) = (mul (inv y) (inv x)) := by 60 | eggxplosion [assocMul, invLeft, mulOne, oneMul, invRight] (timeLimit := 5) 61 | 62 | #print inv_mul 63 | 64 | theorem one_inv 65 | {G: Type} 66 | (inv: G → G) 67 | (mul: G → G → G) 68 | (one: G) 69 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 70 | (invLeft: forall (a: G), mul (inv a) a = one) 71 | (mulOne: forall (a: G), a = mul a one) 72 | (oneMul: forall (a: G), mul one a = a) 73 | (invRight: forall (a: G), one = mul a (inv a)) 74 | (x y : G) 75 | : (inv one) = one := by 76 | eggxplosion [assocMul, invLeft, mulOne, oneMul, invRight] 77 | 78 | #print inv_mul_cancel_left 79 | 80 | /- 81 | theorem test_simplification 82 | {G: Type} 83 | (inv: G → G) 84 | (mul: G → G → G) 85 | (one: G) 86 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 87 | (invLeft: forall (a: G), mul (inv a) a = one) 88 | (mulOne: forall (a: G), a = mul a one) 89 | (oneMul: forall (a: G), mul one a = a) 90 | (invRight: forall (a: G), one = mul a (inv a)) 91 | (x y : G) 92 | -- not equal terms, but simplies 93 | : (inv (mul x (inv y))) = one := by 94 | eggxplosion [assocMul, invLeft, mulOne, oneMul, invRight] 95 | 96 | 97 | #print test_simplification 98 | -/ 99 | -------------------------------------------------------------------------------- /Evaluation/GroupsKnuthBendix.v: -------------------------------------------------------------------------------- 1 | Theorem inv_inv 2 | {G: Type} 3 | (inv: G -> G) 4 | (mul: G -> G -> G) 5 | (one: G) 6 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 7 | (invLeft: forall (a: G), mul (inv a) a = one) 8 | (mulOne: forall (a: G), a = mul a one) 9 | (oneMul: forall (a: G), mul one a = a) 10 | (invRight: forall (a: G), one = mul a (inv a)) 11 | (x: G) 12 | : (inv (inv x) = x). 13 | Proof. 14 | try congruence with assocMul invLeft mulOne oneMul invRight. 15 | Admitted. 16 | 17 | Theorem inv_mul_cancel_left 18 | {G: Type} 19 | (inv: G -> G) 20 | (mul: G -> G -> G) 21 | (one: G) 22 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 23 | (invLeft: forall (a: G), mul (inv a) a = one) 24 | (mulOne: forall (a: G), a = mul a one) 25 | (oneMul: forall (a: G), mul one a = a) 26 | (invRight: forall (a: G), one = mul a (inv a)) 27 | (x y : G) 28 | : (mul (inv x) (mul x y)) = y. 29 | Proof. 30 | try congruence with assocMul invLeft mulOne oneMul invRight. 31 | Qed. 32 | 33 | Theorem mul_inv_cancel_left 34 | {G: Type} 35 | (inv: G -> G) 36 | (mul: G -> G -> G) 37 | (one: G) 38 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 39 | (invLeft: forall (a: G), mul (inv a) a = one) 40 | (mulOne: forall (a: G), a = mul a one) 41 | (oneMul: forall (a: G), mul one a = a) 42 | (invRight: forall (a: G), one = mul a (inv a)) 43 | (x y : G) 44 | : (mul x (mul (inv x) y)) = y. 45 | Proof. 46 | try congruence with assocMul invLeft mulOne oneMul invRight. 47 | Qed. 48 | 49 | Theorem inv_mul 50 | {G: Type} 51 | (inv: G -> G) 52 | (mul: G -> G -> G) 53 | (one: G) 54 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 55 | (invLeft: forall (a: G), mul (inv a) a = one) 56 | (mulOne: forall (a: G), a = mul a one) 57 | (oneMul: forall (a: G), mul one a = a) 58 | (invRight: forall (a: G), one = mul a (inv a)) 59 | (x y : G) 60 | : (inv (mul x y)) = (mul (inv y) (inv x)). 61 | Proof. 62 | try congruence with assocMul invLeft mulOne oneMul invRight. 63 | Admitted. 64 | 65 | Theorem one_inv 66 | {G: Type} 67 | (inv: G -> G) 68 | (mul: G -> G -> G) 69 | (one: G) 70 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 71 | (invLeft: forall (a: G), mul (inv a) a = one) 72 | (mulOne: forall (a: G), a = mul a one) 73 | (oneMul: forall (a: G), mul one a = a) 74 | (invRight: forall (a: G), one = mul a (inv a)) 75 | (x y : G) 76 | : (inv one) = one. 77 | Proof. 78 | try congruence with assocMul invLeft mulOne oneMul invRight. 79 | Qed. 80 | -------------------------------------------------------------------------------- /Evaluation/GroupsKnuthBendixAesop.lean: -------------------------------------------------------------------------------- 1 | -- this only works with the nightly 2022-07-13, adding results as comments 2 | import Aesop 3 | 4 | theorem inv_inv 5 | {G: Type} 6 | (inv: G → G) 7 | (mul: G → G → G) 8 | (one: G) 9 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 10 | (invLeft: forall (a: G), mul (inv a) a = one) 11 | (mulOne: forall (a: G), mul a one = a) 12 | (oneMul: forall (a: G), mul one a = a) 13 | (invRight: forall (a: G), mul a (inv a) = one) 14 | (x: G) 15 | : (inv (inv x) = x) := by 16 | aesop -- aesop: failed to prove the goal after exhaustive search. 17 | 18 | theorem inv_mul_cancel_left 19 | {G: Type} 20 | (inv: G → G) 21 | (mul: G → G → G) 22 | (one: G) 23 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 24 | (invLeft: forall (a: G), mul (inv a) a = one) 25 | (mulOne: forall (a: G), mul a one = a) 26 | (oneMul: forall (a: G), mul one a = a) 27 | (invRight: forall (a: G), mul a (inv a) = one) 28 | (x y : G) 29 | : (mul (inv x) (mul x y)) = y := by 30 | aesop -- goals accomplished 31 | 32 | 33 | theorem mul_inv_cancel_left 34 | {G: Type} 35 | (inv: G → G) 36 | (mul: G → G → G) 37 | (one: G) 38 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 39 | (invLeft: forall (a: G), mul (inv a) a = one) 40 | (mulOne: forall (a: G), mul a one = a) 41 | (oneMul: forall (a: G), mul one a = a) 42 | (invRight: forall (a: G), mul a (inv a) = one) 43 | (x y : G) 44 | : (mul x (mul (inv x) y)) = y := by 45 | aesop -- goals accomplished 46 | 47 | theorem inv_mul 48 | {G: Type} 49 | (inv: G → G) 50 | (mul: G → G → G) 51 | (one: G) 52 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 53 | (invLeft: forall (a: G), mul (inv a) a = one) 54 | (mulOne: forall (a: G), mul a one = a) 55 | (oneMul: forall (a: G), mul one a = a) 56 | (invRight: forall (a: G), mul a (inv a) = one) 57 | (x y : G) 58 | : (inv (mul x y)) = (mul (inv y) (inv x)) := by 59 | aesop -- aesop: failed to prove the goal after exhaustive search. 60 | 61 | theorem one_inv 62 | {G: Type} 63 | (inv: G → G) 64 | (mul: G → G → G) 65 | (one: G) 66 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 67 | (invLeft: forall (a: G), mul (inv a) a = one) 68 | (mulOne: forall (a: G), mul a one = a) 69 | (oneMul: forall (a: G), mul one a = a) 70 | (invRight: forall (a: G), mul a (inv a) = one) 71 | (x y : G) 72 | : (inv one) = one := by 73 | aesop -- aesop: failed to prove the goal after exhaustive search. 74 | -------------------------------------------------------------------------------- /Evaluation/GroupsKnuthBendixAuto.v: -------------------------------------------------------------------------------- 1 | Parameter G : Type. 2 | Parameter inv: G -> G. 3 | Parameter mul: G -> G -> G. 4 | Parameter one: G. 5 | Axiom assocMul : forall (a b c: G), mul a (mul b c) = (mul (mul a b) c). 6 | Axiom invLeft: forall (a: G), mul (inv a) a = one. 7 | Axiom oneMul: forall (a: G), mul one a = a. 8 | Axiom mulOne: forall (a: G), mul a one = a. 9 | Axiom invRight: forall (a: G), mul a (inv a) = one. 10 | Global Hint Rewrite assocMul invLeft mulOne oneMul invRight : base0. 11 | 12 | Theorem inv_mul_cancel_left 13 | (x y : G) 14 | : (mul (inv x) (mul x y)) = y. 15 | Proof. 16 | try autorewrite with base0. reflexivity. Qed. 17 | 18 | Theorem mul_inv_cancel_left 19 | (x y : G) 20 | : (mul x (mul (inv x) y)) = y. 21 | Proof. 22 | try autorewrite with base0. reflexivity. Qed. 23 | 24 | Theorem inv_mul 25 | (x y : G) 26 | : (inv (mul x y)) = (mul (inv y) (inv x)). 27 | Proof. 28 | try autorewrite with base0. Admitted. 29 | 30 | Theorem one_inv 31 | (x y : G) 32 | : (inv one) = one. 33 | Proof. 34 | try autorewrite with base0. Admitted. 35 | 36 | Theorem inv_inv 37 | (x: G) 38 | : (inv (inv x) = x). 39 | Proof. 40 | try autorewrite with base0. Admitted. 41 | -------------------------------------------------------------------------------- /Evaluation/GroupsKnuthBendixExploded.v: -------------------------------------------------------------------------------- 1 | Theorem inv_inv 2 | {G: Type} 3 | (inv: G -> G) 4 | (mul: G -> G -> G) 5 | (one: G) 6 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 7 | (assocMul': forall (a b c: G), (mul (mul a b) c) = mul a (mul b c)) 8 | (invLeft: forall (a: G), mul (inv a) a = one) 9 | (invLeft': forall (a: G), one = mul (inv a) a) 10 | (mulOne: forall (a: G), a = mul a one) 11 | (mulOne': forall (a: G), mul a one = a) 12 | (oneMul: forall (a: G), mul one a = a) 13 | (oneMul': forall (a: G), a = mul one a) 14 | (invRight: forall (a: G), one = mul a (inv a)) 15 | (invRight': forall (a: G), mul a (inv a) = one) 16 | (x: G) 17 | (invRightx : mul x (inv x) = one) 18 | (invRightOne : mul one (inv one) = one) 19 | (invRightx' : one = mul x (inv x)) 20 | (invRightOne' : one = mul one (inv one)) 21 | : (inv (inv x) = x). 22 | Proof. 23 | try congruence with assocMul assocMul' invLeft invLeft' mulOne mulOne' oneMul oneMul' invRight invRightx invRightx' invRightOne invRightOne'. 24 | Qed. 25 | 26 | Theorem inv_mul_cancel_left 27 | {G: Type} 28 | (inv: G -> G) 29 | (mul: G -> G -> G) 30 | (one: G) 31 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 32 | (invLeft: forall (a: G), mul (inv a) a = one) 33 | (mulOne: forall (a: G), a = mul a one) 34 | (oneMul: forall (a: G), mul one a = a) 35 | (invRight: forall (a: G), one = mul a (inv a)) 36 | (x y : G) 37 | : (mul (inv x) (mul x y)) = y. 38 | Proof. 39 | try congruence with assocMul invLeft mulOne oneMul invRight. 40 | Qed. 41 | 42 | Theorem mul_inv_cancel_left 43 | {G: Type} 44 | (inv: G -> G) 45 | (mul: G -> G -> G) 46 | (one: G) 47 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 48 | (invLeft: forall (a: G), mul (inv a) a = one) 49 | (mulOne: forall (a: G), a = mul a one) 50 | (oneMul: forall (a: G), mul one a = a) 51 | (invRight: forall (a: G), one = mul a (inv a)) 52 | (x y : G) 53 | : (mul x (mul (inv x) y)) = y. 54 | Proof. 55 | try congruence with assocMul invLeft mulOne oneMul invRight. 56 | Qed. 57 | 58 | Theorem inv_mul 59 | {G: Type} 60 | (inv: G -> G) 61 | (mul: G -> G -> G) 62 | (one: G) 63 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 64 | (invLeft: forall (a: G), mul (inv a) a = one) 65 | (mulOne: forall (a: G), a = mul a one) 66 | (oneMul: forall (a: G), mul one a = a) 67 | (invRight: forall (a: G), one = mul a (inv a)) 68 | (x y : G) 69 | (assocMul': forall (a b c: G), (mul (mul a b) c) = mul a (mul b c)) 70 | (invLeft': forall (a: G), one = mul (inv a) a) 71 | (mulOne': forall (a: G), mul a one = a) 72 | (oneMul': forall (a: G), a = mul one a) 73 | (invRight': forall (a: G), mul a (inv a) = one) 74 | (invRightx : mul x (inv x) = one) 75 | (invRightOne : mul one (inv one) = one) 76 | (invRightx' : one = mul x (inv x)) 77 | (invRightOne' : one = mul one (inv one)) 78 | (invRighty : mul y (inv y) = one) 79 | (invRighty' : one = mul y (inv y)) 80 | : (inv (mul x y)) = (mul (inv y) (inv x)). 81 | Proof. 82 | try congruence with assocMul assocMul' invLeft invLeft' mulOne mulOne' oneMul oneMul' invRight invRightx invRightx' invRightOne invRightOne' invRighty invRighty'. 83 | Admitted. 84 | 85 | Theorem one_inv 86 | {G: Type} 87 | (inv: G -> G) 88 | (mul: G -> G -> G) 89 | (one: G) 90 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 91 | (invLeft: forall (a: G), mul (inv a) a = one) 92 | (mulOne: forall (a: G), a = mul a one) 93 | (oneMul: forall (a: G), mul one a = a) 94 | (invRight: forall (a: G), one = mul a (inv a)) 95 | (x y : G) 96 | : (inv one) = one. 97 | Proof. 98 | try congruence with assocMul invLeft mulOne oneMul invRight. 99 | Qed. 100 | -------------------------------------------------------------------------------- /Evaluation/GroupsKnuthBendixSimp.lean: -------------------------------------------------------------------------------- 1 | theorem inv_inv 2 | {G: Type} 3 | (inv: G → G) 4 | (mul: G → G → G) 5 | (one: G) 6 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 7 | (invLeft: forall (a: G), mul (inv a) a = one) 8 | (mulOne: forall (a: G), mul a one = a) 9 | (oneMul: forall (a: G), mul one a = a) 10 | (invRight: forall (a: G), mul a (inv a) = one) 11 | (x: G) 12 | : (inv (inv x) = x) := by 13 | simp [assocMul, invLeft, mulOne, oneMul, invRight] 14 | 15 | theorem inv_mul_cancel_left 16 | {G: Type} 17 | (inv: G → G) 18 | (mul: G → G → G) 19 | (one: G) 20 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 21 | (invLeft: forall (a: G), mul (inv a) a = one) 22 | (mulOne: forall (a: G), mul a one = a) 23 | (oneMul: forall (a: G), mul one a = a) 24 | (invRight: forall (a: G), mul a (inv a) = one) 25 | (x y : G) 26 | : (mul (inv x) (mul x y)) = y := by 27 | simp [assocMul, invLeft, mulOne, oneMul, invRight] 28 | 29 | 30 | theorem mul_inv_cancel_left 31 | {G: Type} 32 | (inv: G → G) 33 | (mul: G → G → G) 34 | (one: G) 35 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 36 | (invLeft: forall (a: G), mul (inv a) a = one) 37 | (mulOne: forall (a: G), mul a one = a) 38 | (oneMul: forall (a: G), mul one a = a) 39 | (invRight: forall (a: G), mul a (inv a) = one) 40 | (x y : G) 41 | : (mul x (mul (inv x) y)) = y := by 42 | simp [assocMul, invLeft, mulOne, oneMul, invRight] 43 | 44 | 45 | theorem inv_mul 46 | {G: Type} 47 | (inv: G → G) 48 | (mul: G → G → G) 49 | (one: G) 50 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 51 | (invLeft: forall (a: G), mul (inv a) a = one) 52 | (mulOne: forall (a: G), mul a one = a) 53 | (oneMul: forall (a: G), mul one a = a) 54 | (invRight: forall (a: G), mul a (inv a) = one) 55 | (x y : G) 56 | : (inv (mul x y)) = (mul (inv y) (inv x)) := by 57 | simp [assocMul, invLeft, mulOne, oneMul, invRight] 58 | 59 | theorem one_inv 60 | {G: Type} 61 | (inv: G → G) 62 | (mul: G → G → G) 63 | (one: G) 64 | (assocMul: forall (a b c: G), mul a (mul b c) = (mul (mul a b) c)) 65 | (invLeft: forall (a: G), mul (inv a) a = one) 66 | (mulOne: forall (a: G), mul a one = a) 67 | (oneMul: forall (a: G), mul one a = a) 68 | (invRight: forall (a: G), mul a (inv a) = one) 69 | (x y : G) 70 | : (inv one) = one := by 71 | simp [assocMul, invLeft, mulOne, oneMul, invRight] 72 | -------------------------------------------------------------------------------- /Evaluation/Reviewer1/counts.sh: -------------------------------------------------------------------------------- 1 | #!/usr/bin/env bash 2 | 3 | set -eu 4 | 5 | header() { 6 | echo "rewrites|proof|time" 7 | } 8 | 9 | gen_count() { 10 | python3 ./gen_count.py "$@" 11 | } 12 | 13 | warmup() { 14 | for proof in autorewrite congruence simp; do 15 | gen_count --proof $proof --time 300 >/dev/null 16 | done 17 | } 18 | 19 | gen_data() { 20 | for proof in autorewrite congruence simp; do 21 | for n in $(seq 1 15) $(seq 16 2 24) $(seq 25 3 60); do 22 | out=$(gen_count --proof $proof --time $((n * n))) 23 | echo "$out" 24 | # if proof times out, don't try larger numbers 25 | if [[ "$out" =~ .*\|inf$ ]]; then 26 | break 27 | fi 28 | if [ "$n" -ge 10 ]; then 29 | sleep 2 30 | fi 31 | done 32 | # there are big jumps at powers of 2 33 | for n in $(seq 5 10); do 34 | gen_count --proof $proof --time $((2 ** n - 1)) 35 | gen_count --proof $proof --time $((2 ** n)) 36 | sleep 2 37 | done 38 | done 39 | } 40 | 41 | if [ "$#" -ge 0 ]; then 42 | out="$1" 43 | # avoid clobbering output in case run is cancelled 44 | temp=".$out.tmp" 45 | header | tee "$temp" 46 | warmup 47 | gen_data | tee -a "$temp" 48 | mv "$temp" "$out" 49 | else 50 | header 51 | warmup 52 | gen_data 53 | fi 54 | -------------------------------------------------------------------------------- /Evaluation/Reviewer1/gen_count.py: -------------------------------------------------------------------------------- 1 | #!/usr/bin/env python3 2 | 3 | from os import path 4 | import math 5 | import re 6 | import subprocess 7 | import sys 8 | import tempfile 9 | import time 10 | 11 | 12 | def coq_B(f): 13 | print("Inductive B := I | O.", file=f) 14 | 15 | 16 | def lean_B(f): 17 | print("inductive B", file=f) 18 | print("| I | O", file=f) 19 | print("open B", file=f) 20 | 21 | 22 | def count_sig(N, f): 23 | args = ["B"] * N 24 | args.append("B") # return type 25 | arg_s = " -> ".join(args) 26 | print(f"(count: {arg_s})", file=f) 27 | 28 | 29 | def count_lemma(N, n, f): 30 | free_vars = [] 31 | for i in reversed(range(n + 1, N + 1)): 32 | free_vars.append(f"b{i}") 33 | if free_vars: 34 | params = "forall {},".format(" ".join(free_vars)) 35 | else: 36 | params = "" 37 | args_l = " ".join(free_vars) 38 | args_r = " ".join(free_vars) 39 | # we have n arguments left 40 | args_l += " O" + " I" * (n - 1) 41 | args_r += " I" + " O" * (n - 1) 42 | print(f"(count_{n} : {params}", file=f) 43 | print(f" count {args_l} = count {args_r})", file=f) 44 | 45 | 46 | def hint_rewrite(N, f): 47 | count_lemmas = [f"count_{n}" for n in range(1, N + 1)] 48 | print("#[local] Hint Rewrite {} : count.".format(" ".join(count_lemmas)), file=f) 49 | 50 | 51 | def theorem_statement(rewrites, N, f): 52 | # strip the leading "0b" from the binary representation 53 | bin_num = bin(rewrites)[2:] 54 | if len(bin_num) < N: 55 | bin_num = "0" * (N - len(bin_num)) + bin_num 56 | count_args = bin_num.replace("0", "O ").replace("1", "I ") 57 | print(" count {} = count {}".format(count_args, "O " * N), file=f, end="") 58 | 59 | 60 | def coq_theorem(rewrites, N, f): 61 | print(f"(* {rewrites} rewrites *)", file=f) 62 | print("Theorem count_upward :", file=f) 63 | theorem_statement(rewrites, N, f) 64 | print(".", file=f) 65 | 66 | 67 | def coq_proof(rewrites, N, proof, f): 68 | print("Proof.", file=f) 69 | print(f' idtac "rewrites = {rewrites}".', file=f) 70 | if proof == "congruence": 71 | for n in range(1, N + 1): 72 | print(f"pose proof count_{n}.", file=f) 73 | # This number is the number of quantified hypotheses that can be added 74 | # to the goal. The default is 1000, which we increase here. Without this 75 | # parameter congruence would simply fail at some earlier point, but we 76 | # want to measure its performance a bit further. 77 | print('time "congruence" congruence 2000.', file=f) 78 | elif proof == "autorewrite": 79 | print(' time "autorewrite" autorewrite with count. reflexivity.', file=f) 80 | print("Qed.", file=f) 81 | 82 | 83 | def lean_proof(N, f): 84 | print("begin", file=f) 85 | count_lemmas = ", ".join([f"count_{n}" for n in range(1, N + 1)]) 86 | print(f" simp [{count_lemmas}],", file=f) 87 | print("end", file=f) 88 | 89 | 90 | def num_digits(rewrites): 91 | """Compute the number of digits required to represent a number of 92 | rewrites.""" 93 | return int(math.floor(math.log(rewrites, 2))) + 1 94 | 95 | 96 | def coq_file(args, f): 97 | rewrites = args.rewrites 98 | N = num_digits(rewrites) 99 | coq_B(f) 100 | print("Context", file=f) 101 | print(f"(* N = {N} *)", file=f) 102 | count_sig(N, f) 103 | for n in range(1, N + 1): 104 | count_lemma(N, n, f) 105 | print(".", file=f) 106 | print(file=f) 107 | hint_rewrite(N, f) 108 | print(file=f) 109 | coq_theorem(rewrites, N, f) 110 | coq_proof(rewrites, N, args.proof, f) 111 | 112 | 113 | def lean_file(args, f): 114 | rewrites = args.rewrites 115 | N = num_digits(rewrites) 116 | lean_B(f) 117 | print(file=f) 118 | print("theorem count_upward", file=f) 119 | count_sig(N, f) 120 | for n in range(1, N + 1): 121 | count_lemma(N, n, f) 122 | print(f" -- rewrites = {rewrites}", file=f) 123 | print(" :", end="", file=f) 124 | theorem_statement(rewrites, N, f) 125 | print(":=", file=f) 126 | lean_proof(N, f) 127 | 128 | 129 | def run_coqc(f): 130 | out = subprocess.run(["coqc", f], capture_output=True) 131 | if out.returncode != 0: 132 | if re.search("Tactic failure", out.stderr.decode("utf-8")): 133 | return None 134 | else: 135 | # something else seems to have gone wrong 136 | print(out.stderr, file=sys.stderr) 137 | print("Coq failed", file=sys.stderr) 138 | sys.exit(1) 139 | return out.stdout.decode("utf-8") 140 | 141 | 142 | def time_coqc(args, time_complete=True): 143 | """Time running coqc on a file generated from args. 144 | 145 | If time_complete is True, then time the entire run rather than just the 146 | autorewrite/congruence invocation. This includes the kernel checking the 147 | resulting proof. 148 | """ 149 | with tempfile.TemporaryDirectory() as tempd: 150 | fname = path.join(tempd, "count.v") 151 | with open(fname, "w") as f: 152 | coq_file(args, f) 153 | start = time.time() 154 | output = run_coqc(fname) 155 | if output is None: 156 | return float("inf") 157 | elapsed = time.time() - start 158 | if not time_complete: 159 | # replace elapsed with the tactic timing result (from the Coq output) 160 | m = re.search(r"""ran for (?P