The Automation of Inductive Theorem Proving

  T H E U S I T Y O F E D I U R G H N B

  IJCAR-04 Tutorial 1 The Automation of Inductive

  Theorem Proving Alan Bundy University of Edinburgh The Need for Inductive Inference Proof by mathematical induction required for reasoning 2 about repetition , e.g. in: • recursive datatypes: numbers, lists, sets, trees, etc;

  • iterative or recursive computer programs; • electronic circuits with loops or parameterisation.

  So needed for proof obligations in formal methods .

  Recursive Datatypes Examples: natural numbers, lists, sets, trees.

  Made From: base and step constructor functions . Naturals: 0:nat and s : nat 7→ nat. 3 nat nil

  Lists: :list(τ ) and :: : τ × list(τ) 7→ list(τ). list :: nil, α :: α

  1

  2

  1 (τ ) = {nil, α :: nil, . . .} where α :τ .

  non-free: some non-identical terms are equal.

  Recursive Definitions 1 Add Two Numbers:

  0 + N = N s(M ) + N = s(M + N ) 4

  • Even Numbers:

  even(0) ¬even(s(0)) even(s(s(N ))) ↔ even(N ) Append Two Lists: nil <> L = L (H :: T ) <> L = H :: (T <> L) Length of a List: len(nil) = 0 len(H :: T ) = s(len(T ))

  Reverse a List: rev(nil) = nil 5 rev(H :: T ) = rev(T ) <> (H :: nil)

  (Quickly) Reverse a List: (tail recursive)

qrev(nil, L) = L

qrev(H :: T, L) = qrev(T, H :: L)

  Rotate a List: rot(0, L) = L rot(s(N ), nil) = nil rot(s(N ), H :: T ) = rot(N, T <> (H :: nil))

  Some Induction Rules Natural Numbers – One-Step:

  P (0), ∀n:nat. (P (n) → P (s(n))) ∀n:nat. P (n) 6 Natural Numbers – Two-Step:

  P (0), P (s(0)), ∀n:nat. (P (n) → P (s(s(n)))) ∀n:nat. P (n)

Lists – One-Step:

  P (nil) ∀h:τ.∀t:list(τ). P (t) → P (h :: t) ∀l:list(τ). P (l)

  Additional Universal Variables Induction Formula:

  ∀n:nat. ∀m:nat. Q(n, m) Step Case: 7 ∀n:nat. ( ∀m:nat. Q(n, m) →

  ) ∀m:nat. Q(s(n), m)

  With Quantifiers Stripped:

Q (n, M ) ⊢ Q(s(n), m)

  NB – m becomes free variable in hypothesis; constant in conclusion! This gives important freedom in inductive proof.

  Theoretical Limitations of Inductive Inference Incompleteness: (G¨odel) formal theories exclude some theorems , 8 examples found by Kirby & Paris.

  Undecidability of Halting Problem: (Turing) no algorithmic test for termination .

  Failure of Cut Elimination: (Kreisel). need to generalise and introduce lemmas .

  (Lack of) Cut Elimination Gentzen’s Cut Rule:

  A , A Γ ⊢ ∆, Γ ⊢ Γ ⊢ ∆ 9 lacks subformula property . Cut Elimination Theorem: Gentzen showed Cut Rule redundant in FOL.

  Kreisel showed necessary in inductive theories .

  Practical Consequences: Need to generalise conjectures.

  Need to introduce lemmas .

  Need for Intermediate Lemmas Conjecture:

  ∀l:list(τ). rev(rev(l)) = l Rewrite Rules: 1 rev (nil) ⇒ nil rev (H :: T ) ⇒ rev(T ) <> (H :: nil)

  Step Case: rev rev

  (h :: t) ) = h :: t (rev(t)) = t ⊢ rev( rev (rev(t) <> (h :: nil)) = h :: t

  ⊢ | {z } blocked Introducing an Intermediate Lemma Lemma Required: rev (X <> Y ) ⇒ rev(Y ) <> rev(X) 1 Cut Rule: introduces this:

  Original: Γ ⊢ rev(rev(l)) = l

  New: rev

  Γ, (X <> Y ) ⇒ rev(Y ) <> rev(X) ⊢ rev(rev(l)) = l rev

  Justification: Γ ⊢ (X <> Y ) ⇒ rev(Y ) <> rev(X) Heuristics needed: to speculate lemma.

Step Case Unblocked

  rev(rev(t)) = t ⊢ rev(rev(t) <> (h :: nil)) = h :: t 1 ⊢ rev(h :: nil) <> rev(rev(t)) = h :: t ⊢ ( rev(nil) <> (h :: nil)) <> rev(rev(t)) = h :: t ⊢ ( nil <> (h :: nil) ) <> rev(rev(t)) = h :: t ⊢ (h :: nil) <> rev(rev(t)) = h :: t ⊢ h :: ( nil <> rev(rev(t)) ) = h :: t ⊢ h :: rev(rev(t)) = h :: t ⊢ h = h ∧ rev(rev(t)) = t fertilization now possible. Prove the following lemma by induction. rev(X <> Y ) = rev(Y ) <> rev(X)

  1. What induction variable and induction rule will you use?

  2. State the base case(s) and give its(their) proof(s)

  3. State the step case(s) and give its(their) proof(s) 1 Assume the axioms of equality, the following recursive definitions of <> and rev, plus the following two further lemmas. nil <> L = L (H :: T ) <> L = H :: (T <> L) rev(nil) = nil rev(H :: T ) = rev(T ) <> (H :: nil)

  L <> nil = L X <> (Y <> Z) = (X <> Y ) <> Z

Solution to Exercise 1 1. The proof works best by induction on X and one-step induction on lists

  2. Base Case: rev(nil <> y) = rev(y) <> rev(nil) rev(y) = rev(y) <> nil 1 rev(y) = rev(y)

  3. Step Case: rev(t <> Y ) = rev(Y ) <> rev(t) ⊢ rev((h :: t) <> y) = rev(y) <> rev(h :: t) ⊢ rev(h :: (t <> y)) = rev(y) <> (rev(t) <> (h :: nil)) ⊢ rev(t <> y) <> (h :: nil) = (rev(y) <> rev(t)) <> (h :: nil) ⊢ rev(t <> y) = rev(y) <> rev(t)

  Reverse Example Revisited Blocked Step Case: rev(rev(t)) = t ⊢ rev(rev(t) <> (h :: nil)) = h :: t

  | {z } blocked

  Weak Fertilization: 1 ⊢ rev(rev(t) <> (h :: nil)) = h :: rev(rev(t)) New Step Case: rev(rev( t ) <> (h :: nil)) = h :: rev(rev( t )) ′ ′ ′ ′ ⊢ rev( rev(h :: t ) <> (h :: nil)) = h :: rev( rev(h :: t ) ) ′ ′ ⊢ rev(( rev(t ) <> (h :: nil) ) <> (h :: nil))

  

| {z }

blocked

′ ′

= h :: rev( rev(t ) <> (h :: nil) )

  | {z } blocked Generalising a Sub-Term Generalisation Required: ′′ ′′ rev( t <> (h :: nil)) = h :: rev( t ) 1 ′′ Step Case Unblocked:

′′

rev(t <> (h :: nil)) = h :: rev(t ) ⊢ ′′ ′′ ′′ ′′ rev( (h :: t ) <> (h :: nil) ) = h :: rev(h :: t ) ′′ ′′ ′′ ′′ rev(h :: (t <> (h :: nil))) = h :: (rev(t ) <> (h :: nil)) ′′ ′′ ′′ ′′ rev(t <> (h :: nil)) <> (h :: nil) = (h :: rev(t )) <> (h :: nil) ′′ ′′ ′′ ′′ rev(t <> (h :: nil)) = h :: rev(t ) ∧ h :: nil = h :: nil

  NB – <> definition applied right to left. Induction Strategy Induction Base

  Step Ripple Symbolic

  Wave 1 Evaluation Fertilization

  Preconditions Declarative: Rippling must be possible in step cases.

  Procedural: Look-ahead to choose induction rule that will permit rippling.

  Annotated Step Case t <> (Y <> Z) = (t <> Y ) <> Z

  ↑ ↑ t <> h t <> y (y <> z) = ( :: ) <> z

  ⊢ h :: ↑ ↑ t <> t <> y <> z h

  (y <> z) = :: 1 ⊢ h ::

  ↑ t <> h (y <> z) = :: (t <> y) <> z

  ⊢ h :: ↑ t <> (y <> z) = (t <> y) <> z

  ⊢ h = h ∧ ↑ (wave-fronts).

  • changing bits in orange boxes

  red (skeleton).

  • unchanging bits in

  Annotated Rewrite Rules ↑

  ↑ H T <> L T <> L ::

  ⇒ H :: ↑ ↑

  ↑

  X Y

  X X :: = Y :: = Y = Y

  1

  2

  1

  2

  1

  1

  2

  2 1 ⇒ X ∧ wave-rules . outward movement of wave-fronts.

  • Note • We reason backwards, so implication direction inverse to rewrite direction, e.g.,

X = Y = Y :: X = Y :: Y

  1

  1

  2

  2

  1

  2

  1

  2

∧ X → X Another View of Rippling ↑

  <> <> h <> :: ☞☞ ▲▲ ✡ ✡ ❏ ❏ ☞☞ ▲▲

  <> z h :: <> z <> z

  • 2 ✔ ✔ ❚ ❚ ✂✂ ❇❇

  ✂✂ ❇❇ h :: t y t y t y ↑ h t <> y h t <> y <> z h

  ( :: ) <> z :: :: (t <> y) <> z Before During After

  Exhaustive Rewriting (without Rippling) Rewrite Rule:

  (X + Y ) + Z ⇒ X + (Y + Z) a 2 Given: + b = 42 Goal: ((c + d) + a) + b = (c + d) + 42 Rewritings: ((c + d) + a) + b = c + (d + 42)

  × (c + (d + a)) + b = (c + d) + 42

  × √

  (c + d) + (a + b) = (c + d) + 42

  Selective Rewriting (with Rippling) ↑

  ↑ Y

  X Y Wave Rule: (X + ) + Z ⇒ + ( + Z ) a

  Given: + b = 42 ↑

  ↑ 2 a

  Goal: ((c + d) + ) + b = (c + d) +

  42 ↑ ↑ a

  Ripple: (c + d) + ( + b ) = (c + d) +

  42 Non-Ripples:

  ↑ a c

  ((c + d) + )) + b = + (d + 42 )

  ↑ a c + (d + ) + b = (c + d) +

  42 Bi-Directional Rippling Wave-Rules:

  T <> L T <> L (H :: ) ⇒ H :: ( ) ↑ ↑

  H T H :: ( <> L) ⇒ ( :: T ) <> L

  • 2 Rippling: –

  ′

rev <> (t (h :: nil)) = h :: rev(t )

  ′ ′ rev t <> t

  ⊢ ( (h :: ) (h :: nil) ) = h :: rev( h :: )

  ′ ⊢ rev ( h :: ( t <> (h :: nil) ) ) = h :: ( rev (t ) <> (h :: nil))

  ′ ′ rev <> h

  ⊢ (t (h :: nil)) <> (h :: nil) = ( :: rev(t ) ) <> (h :: nil) ′ ′

  ′ rev <>

  ⊢ (t (h :: nil)) = h :: rev(t ) ∧ h :: nil = h :: nil

  Rippling Sideways and In rot (len(t), t <> K) = K <> t ↑ ↑ rot t t <> t

  ⊢ (len( h :: ), h :: ⌊k⌋) = ⌊k⌋ <> h :: ↑ ↑ rot len , t <> t

  • 2 ⊢ ( s( (t) ) h :: ⌊k⌋ ) = ⌊k⌋ <> h ::
  • rot t rot

      <> t ⊢ (len(t), t <> ⌊k⌋ <> (h :: nil) ) = ⌊k⌋ <> (h :: nil) rot

      ⊢ (len(t), t <> ⌊k <> (h :: nil)⌋) = ⌊k <> (h :: nil)⌋ <> t

    • ⌊ Sinks ⌋ provide alternative wave-front destination, available when free variables are in hypothesis. • Wave-fronts have directions: out / in .

      Sideways and Inwards Wave-rules ↑ ↑

      ↓ rot s N , H T rot <> ( ( ) :: ) (N, T (H :: nil) ) 2

      ⇒ ↑

      ↓ L <> H T L <> <> T ( :: ) (H :: nil)

      ⇒ ↓

      ↓ X <> Y X <>

      ( ) <> Z ( Y <> Z ) ⇒

      Sideways Rippling: Another View φ φ φ φ η η η η 2 ✔ ✔ ❚ ❚ ✔ ✔ ❚ ❚ ✔ ✔ ❚ ❚ ✔ ✔ ❚ ❚ 2 3 ↓ – 1 µ ν c ( µ ) ν µ c ( ν ) µ ν 4 c ( x ) y x y x y x c ( y ) Before Stage-1 Stage-2 After

      The Preconditions of Rippling

      1. The induction conclusion contains a wave-front, ↑ ↑ e.g. rot len , t <> ( s( (t) ) h :: ⌊k⌋ ) = . . . .

      2. A wave-rule applies to this wave-front. ↑ ↑ e.g. rot N , T . . . 2 ( s( ) H :: ) ⇒ .

      3. Any condition is provable. e.g.

      X T

      X X 6= H → ∈ H :: ⇒ ∈ T .

      4. Inserted inwards wave-fronts contain a sink or an outwards wave-front in their wave-hole. e.g. rot (len(t), t <> ⌊k⌋ <> (h :: nil) ) = . . . .

      Alternate Annotations of Wave-Rules ↑

      ↑ X <> <> Z X <> Y ( Y ) ) <> Z

      ⇒ ( ↑ ↓

      X <> Z X <> Z Y <> <> Y

      ( ) ) ⇒ (

      ↓ ↓ X <> Y <> Z X <> Y ( ) ( ) <> Z 2

      ⇒ ↑

      ↑ Y

      Y <> Z X <>

      ( ) <> Z ) ⇒ X <> (

      ↑ ↓

      X <> Y <> Z X <> Y <> Z ( )

      ( ) ⇒

      

      X <> Y X <> ( ) <> Z ( Y <> Z ) ⇒

      Well Annotated Terms Normal form: one functor thick, e.g.

      ↑ ↑ ↑

    • s n n (s( )) ) )
    • 2 7→ s( s( Absorption into sinks: of wave-fronts, e.g.

        ↓ f (a, ⌊b⌋ ) ⌊f (a, b)⌋ 7→

        Absorption of nested annotation: not in wave-holes, e.g.

        ↑ ↑ ↑ f a , b b ( g( ) ) )

        7→ f(g(a), Informal Definition of Wave-Rule Matching LHS = term 3 Matching annotation: There are two conditions:

      1. all annotation in LHS matches annotation in term.

        2. all annotation in term matches annotation or variable in LHS , Matching variables: Variables match as in normal matching.

        Matching term structure: as in normal matching.

        Examples of Matches and Non-Matches LHS of wave-rule Term subst s (X) s (x) do match

        X = x s (X) s ( s ( x ) ) do match ↑ ↑ X = s( x ) s X s s x 3 ( ) ( ( ) ) do not match ↑ – s (

        X ) s (x) do not match s s x (X) + Y ( ) + y do not match ↑ ↑ s ( X + Y ) s ( s( a × b ) + c ) do match X = s( a × b ) ,

        Y = c Skeletons Skeletons: omit wave-fronts,

        ↑ rev t rev e.g. skel ( ( h :: ) (t) ) = { }. e.g.

        Multiple Skeletons: ↑ skel rev

      X <> Y rev , rev

      ( ( ) (X) (Y ) 3

        ) = { } Weakenings: absorb all but one skeleton, ↑ e.g. rev <> Y

        ( X ) Skeleton Preservation: in wave-rules, e.g.

        ↑ ↑ ↑

        X <>

        X Y <> Y

        X X =

        = Y = Y

        1

        2

        1

        2

        1

        1

        2

        2 ⇒ ∧

        X ,

        X = Y = Y

        1

        1

        2

        2 skeleton of both sides: {

        }

      Termination of Rippling Define measure: from annotated terms to well-founded set

        3

      Define wave-rule: to be measure decreasing from LHS to

      RHS, (and skeleton preserving).

        Show measure decreases: when wave-rules applied, because preserved by substitution and subterm replacement.

        Idea of Measure

      • ✡ ❏ ✡ ☎☎ ❉ ↑ ❏ ✡ ❏ ☎ ❉❉

        ↑ z s ) ( s( +

        2

        2 ✡ ❏ ✡

        ☎☎ ❉ 3 ❏ ✡ ❏ ☎ ❉❉

      • s x s y ( ) ( )

        1

        1

        2 ↑ ↑

        ↑ ↑

        s ( s( s( ) s ( ) ) ) + z [2,2,0]

      • x y
      Decrease of Wave Measure ↑ s ( + ) + +

        ☞☞ ▲▲ ✡ ✡ ❏ ❏ ☞☞ ▲▲ ↑ z z z 3

        s ( ) ✡ ✡ ❏ ❏ ✂✂ ❇❇ ✂✂ ❇❇

        ↑ s x y x y x y ( )

        [1,0,0] [0,1,0] [0,0,1] Multi-Hole Rule: ↑ ↑

      • binom

        X , Y binom Y binom ( s( ) s( ) ) ⇒ (X, s( ) ) (X, Y ) Single-Hole Weakenings: ↑ ↑ binom X , Y binom Y 3 ( s( ) s( ) ) ⇒ (X, s( ) ) + binom(X, Y ) ↑ ↑ binom

        X , Y binom ( s( ) s( ) ) ⇒ binom(X, s(Y )) + (X, Y )

        Form Measure of Each Weakening:  

        1   binom (., .) + binom(X, Y ) ⇒ 1 ⇒

        1 ✡ ✡ ❏ ❏ ✄✄ ❈❈

        X Y s( )

        1 General Case 2 Combine Measures of Weakenings: using multi-sets, 3

        out( (X, s( ) ) (X, Y ) ) = {[1, 1], [0, 1]} using multi-set ordering.

      • binom Y binom

        Inwards Measure: similar, but order inverted, only well-founded due to skeleton preservation.

        Overall Measure: lexicographic combination of outwards and inwards.

        Ground Difference Unification Algorithm Mismatching Terms Annotated Term 1 Term 2 Term 1 rev (h :: (t <> l)) rev (t <> l) rev ( h :: (t <> l) ) 3 h :: (t <> l) t <> l h :: (t <> l) (h <> t) <> l h <> l ( h <> t ) <> l

      • s (a) + b s (a) s (a) + b

        (n × m) + m (n × m) + m n × m

        Difference unification algorithm similar regular unification,

      • but with option to hide mis-matching function symbols in wave-fronts.

        Ground difference unification does not instantiate variables.

        Automatic Annotation of Wave-Rules 1 • Given equation, equivalence, etc., e.g.

        X 3

      • (Y + Z) = (X + Y ) + Z • Ground difference unify to get all possible wave annotations.

        X X

      • ( Y + Z ) = ( + Y ) + Z

        X Y Z

      • ( ) = ( X + Y ) + Z + ... ... ...
      Automatic Annotation of Wave-Rules 2 • Calculate measures, e.g. out

        X X ( + ( Y + Z ) ) = [1, 0] out( ( + Y ) + Z ) = [0, 1] out

        X Y Z out

        X ( + + ( ) ) = [1, 0] (( + Y ) + Z) = [1, 0] 4

      • Determine wave-front directions and rule orientation to

      ensure decreasing measure, e.g.

        

        X X

      • ( Y + Z ) + Y ) + Z

        ⇒ ( ↑ ↓

      • X Y Z + ( ) ( X + Y ) + Z

        ⇒ ↑

        ↓

      • X Y Z ( X + Y ) + Z + ( )

        ⇒ ... ... ...

      Advantages of Rippling Selective: not exhaustive rewriting

        4 Bi-directional: rewriting. Termination: of any set of wave-rules, despite bidirectionality.

        

      Heuristic basis: for choosing lemmas, generalisations and

      inductions.

      Exercise 2 –

        4

        1. Annotate the following equation as a wave-rule in as many sum (H :: T ) = s(sum(T )) 2. Give the measure of each side one of your wave-rules.

        1.

      sum T sum

      ( H :: ) ⇒ s( (T ) ) (1) sum sum T s( (T ) ) ⇒ ( H :: ) (2)

      sum T sum

      4

        ( H :: ) ⇒ s( (T ) ) (3) Wave-rule (3) is a bit of an aberration.

        2. Wave Outer Inner Rule LHS RHS LHS RHS (1) [1, 0] [0, 1] [0, 0] [0, 0] (2) [0, 0] [0, 0] [1, 0] [0, 1] (3) [1, 0] [0, 0] [0, 0] [1, 0]

        Ripple-Based Heuristics Induction Rules: choose induction which best 4 supports rippling.

        Lemmas: design wave-rule to unblock ripple.

        Generalisation: generalise goal to allow wave-rule to apply.

      Induction Variable Selection ∀t, l, m:list(nat). –

        4 t <> (l <> m) = (t <> l) <> m ւ ↓ ց t <> (l <> m) = t <> ( l <> m) = t <> (l <> m ) = ( t <> l) <> m (t <> l ) <> m (t <> l) <> m

        Induction Rules Available P (nil) . . . :: . . .

        P P t ∀h:τ.∀t:list(τ ). ( (t) → ( h :: ) ) ∀l:list(τ ). P (l) P (nil) ∀h:τ. P (h :: nil) 4 . . . :: . . . :: . . .

        P P t ∀h , h :τ.∀t:list(τ ). ( (t) → ( h :: h :: ) )

        1

        2 1

        2 ∀l:list(τ ). P (l) P (nil)

        ∀h:τ. P (h :: nil) . . . <> . . . P P P t t

        ∀t , t :list(τ ). ( (t ) ∧ (t ) → ( <> ) )

        1

        2

        1

        2

        1

        2 ∀l:list(τ ). P (l) ... ... Ripple Analysis 1 Wave Rule:

        Y ( X :: ) <> Z ⇒ X :: (Y <> Z)

      • 4 Wave Occurrences: ∀t, l, m:list(nat). t <> ( l <> m) = ( t <> l) <> m Choice of Induction Variable:

        variable occurrences t all unflawed l some unflawed, some flawed m all flawed t is best choice since it has the fewest flaws. Wave Rule: Y ( X :: ) <> Z ⇒ X :: (Y <> Z)

        Induction Conclusion Needed: t t 4 h :: <> (l <> m) = ( h :: <> l) <> m

        Choice of Induction Rule: wave-front induction rule

      . . .

        X :: . . . :: . . .

      . . .

      X :: X :: . . . :: . . . :: . . .

        1

        2 . . . . . . <> . . . <> . . .

        . . .

        X :: suggests . . . :: . . . induction rule. Conjecture: ∀t, l :list(τ ). even(len(t <> l)) ↔ even(len(l <> t)) Wave-Rules: Y

      X ⇒ X

        ( :: ) <> Z :: (Y <> Z) even X even ( (s( )) ) (X )

      s ⇒

      len T len 4
      • len T len

        ( H :: ) ⇒ s ( (T ) ) (L <> H :: ) ⇒ s ( (L <> T ) ) Step Case: using . . . :: . . . induction on t. even t t h h

        (len(( :: :: )) even ) <> l)) ↔ even(len(l <> len

        (len( h :: (t <> l) s ( (l <> t) ) )

      even len len )) ↔ even( ( s ( (t <> l) ) ) s ( (l <> t) ) ) ↔ even( | {z } | {z } blocked blocked because only one-level look-ahead. Failure of Ripple Preconditions

      • Precondition 1 is true : 1. The induction conclusion contains a wave-front.

        even len ( s( (t <> l) ) ) ↔ . . . 5 (other side similar) • Precondition 2 is false : 2. A wave-rule applies to this wave-front. however, there is a near miss (only wave-fronts mismatch): even ( s(s( X )) ) ⇒ even (X) • Preconditions 3 and 4 are inapplicable.

        3. Any condition is provable.

        4. Inserted inwards wave-fronts contain a sink or an outwards wave-front.

        Induction Revision Critic: Invocation Induction

      Ind_Strat 1

      Ind_Strat 2 Base Step 5Blocked Critic Wave Induction Ripple

      Induction Revision Critic: Application

        Ind_Strat 1 Ind_Strat 2 (now abandoned)

      • 5

        Induction Base Step

        

      even s len Blocked Goal: ( ( (t <> l) ) ) ↔ . . .

        Inwards Wave-Rules:

        ⇒ H :: (T <> L) ( H :: T ) <> L s ( len (T ) ) len ( H :: T ) Calculation of New Induction Term:

      even s len

      5 ( (s( (t <> l) )) ) ↔ . . . even s len
      • ( (s( (⌊t⌋ <> l) )) ) ↔ . . .

        

      2

      even ( s ( len ( H :: (⌊t⌋ <> l) ) ) ) ↔ . . .

        1 2 even (len( H :: H :: (⌊t⌋ <> l) )) ↔ . . . 1 even (len( H :: (⌊H :: t⌋ <> l) )) ↔ . . . 2 even (len(⌊H :: H :: t⌋ <> l)) ↔ . . . 1 2 suggests . . . :: . . . :: . . . induction on t. New Induction Rule: 1 2 P (nil), ∀h:τ. P (h :: nil), 1 2 ∀ → h :τ.∀h :τ.∀t:list(τ ). P (t) P ( h :: h :: t )

        ∀ l :list(τ ). P (l) New Step Case: 1 2 1 2 even h t h t 5

        (len(( :: h :: ) <> l)) ↔ even(len(l <> :: h :: ))

        1 2 1 2 even h h t <> l h t

        (len( :: ( :: ) )) ↔ even(len(l <> :: h :: )) 2 2 ↑ even ( s ( len ( h :: (t <> l) ) ) ) ↔ even( s ( len (l <> h :: t ) ) ) even ( s (s( len (t <> l) )) ) ↔ even( s (s( len (l <> t) )) ) even (len(t <> l)) ↔ even(len(l <> t))

        Rippling Failure: Missing Wave-Rule Conjecture:

        ∀n:nat. even(n + n) Wave-Rules: ↑ ↑

        X X 5 s( ) + Y ⇒ s( + Y ) even X even ( s(s( )) ) ⇒ (X)

        Induction Conclusion: ↑ ↑ even n n ( s( ) s( ) ) +

      • even n n ( s( s( ) ) )

        | {z } blocked Failure of Ripple Preconditions

      • Precondition 1 is true : 1. The induction conclusion contains a wave-front.

        ( s( + s( ) ) ) even n n 5

      • – • Precondition 2 is false : 2. A wave-rule applies to this wave-front.

        (to neither of them) • Preconditions 3 and 4 are inapplicable.

        3. Any condition is provable.

        4. Inserted inwards wave-fronts contain a sink or an outwards wave-front.

      Lemma Speculation Critic: Invocation

        Ind_Strat 1 Induction

      • 5 Base

        Step Ripple Wave Lemma Blocked Critic Lemma Speculation Critic: Application Cut Induction Ind_Strat 1 Lemma Base Step 5Unblocked Wave Ripple

      Blocked Goal: even n n ( s( s(

        ) ) ) + focus on innermost wave-front.

        Schematic Wave-Rule: ↑ ↑

      • 5

        X s( Y ) ⇒ F ( X + Y ) Continued Ripple:

      even n

      ( s(F ( + n )) ) even (n + n) where F = s.

        Final Wave-Rule: ↑ ↑

        X s( ) ⇒ s( + Y ) which must now be proved.

      • X Y

      Exercise 3

        Apply the lemma speculation critic to the step case of the proof of: rev(X <> Y ) = rev(Y ) <> rev(X) in the context where only the equality axioms and recursive definitions of 6 <> and rev are available. nil <> L = L

      (H :: T ) <> L = H :: (T <> L)

      rev(nil) = nil rev(H :: T ) = rev(T ) <> (H :: nil)

        Specify the blocked goal, schematic wave-rule, continued ripple and final wave-rule.

        Blocked Goal: rev (t <> y) <> (h :: nil) = rev (y) <> ( rev (t) <> (h :: nil))

        | {z } blocked

        Schematic Wave-Rule: U <>

        V U <> V 6 ( <> W ) ⇒ F ( , W )

        Continued Ripple:

      rev rev (t <> y) <> (h :: nil) = F ( (y) <> rev(t) , h :: nil) rev rev (t <> y) = (y) <> rev(t) where F ≡<>.

        Final Wave-Rule: U <>

        V U <> V ( <> W ) ⇒ <> W Conjecture: ∀t:list(τ). rev(t) = qrev(t, nil)

        Wave-Rules: ↑ ↑

      • rev H T rev <> H ( :: ) (T ) :: nil
      • 6

          ⇒

          ↓ qrev T , L qrev L ( H :: ) (T, H :: )

          ⇒ Induction Conclusion:

          ↑ ↑ rev h t h t , nil ( :: ) = qrev( :: )

          ↑ ↑ rev <> h h t , nil (t) :: nil = qrev( :: )

          | {z } missing sink Failure of Ripple Preconditions 1. The induction conclusion contains a wave-front. . . . = qrev ( h :: t , nil ) 6 2. A wave-rule applies to this wave-front. qrev T , L qrev L ( H :: ) ⇒ (T, H :: )

          Precondition 4 is false : 4. Inserted inwards wave-fronts contain a sink or an outwards wave-front. nil

          . . . = qrev(t, h :: ) Generalisation Critic: Invocation Induction Ind_Strat 1 Ind_Strat 2 isation General- Base Step 6Blocked Critic Wave Gen Ripple

        Generalisation Critic: Application (now abandoned) Ind_Strat 1 Ind_Strat 2 Generalisation Cut
        • 6

          Justification Generalised Conjecture Original Conjecture: ∀t:list(τ). rev(t) = qrev(t, nil)

          Disallowed Ripple: ↓ 6

          . . . h nil = qrev(t, :: )

          Schematic Conjecture: ∀t:list(τ).∀l:list(τ). F (rev(t), l) = qrev(t, G(l))

          Induction Hypothesis: F (rev(t), L) = qrev(t, G(L)) where F , G and L are meta-variables.

          New Step Case: F (rev( h :: t ), ⌊l⌋) = qrev ( h :: t , G (⌊l⌋)) F ( rev(t) <> h :: nil , ⌊l⌋) = qrev (t, h :: G (⌊l⌋) ) 6 rev (t) <> ( h :: nil <> F ( rev(t) <> (h :: nil) , ⌊l⌋) ) = qrev (t, h :: G (⌊l⌋) )

          ′ rev (t) <> ( h :: F ( rev(t) <> (h :: nil) , ⌊l⌋) ) = qrev (t, h :: G (⌊l⌋) ) rev (t) <> (⌊h :: l⌋) = qrev (t, ⌊h :: l⌋) where F =<>, F = λX.λY. Y and G = λX. X.

          X <> Z Key Wave-Rule: ( X <> Y ) <> Z ⇒ ( Y <> ) Generalised Conjecture: ∀t:list(τ ).∀l:list(τ ). rev(t) <> l = qrev(t, l)

        Pattern of Failure Suggests Patch

          Precondition Generalization Case analyse Induction Lemma revision discovery

        • 6

          1 ⋆ ⋆ ⋆ ⋆ 2 ⋆ ⋆

          3

        • ⋆ 4 • ⋆ = success ◦ = partial success • = failure
        No Conjecture (i) (ii) T3 len (X <> Y ) = len(Y ) + len(X) L1 T4 len (X <> X) = double(len(X)) L2 T5 len (rev(X)) = len(X) L3 T8 n th (X, n th (Y, Z)) = n th (Y, n th (X, Z)) L4, L5 T10 rev (rev(X)) = X L8

          T11 rev (rev(X) <> rev(Y )) = Y <> X L9, L10 T12 qrev (X, Y ) = rev(X) <> Y L11 T13 half

          (X + X) = X L1 T14 ordered (isort(X)) L12 T16 even

          (X + X) L1 T21 rotate (len(X), X <> Y ) = Y <> X L11, L13 T22 even

          

        (len(X <> Y )) ↔ even(len(Y <> X))

        ⋆ L14 T23 half (len(X <> Y )) = half (len(Y <> X)) ⋆ L15

          T26 half

        (X + Y ) = half (Y + X)

          L17 (i) induction revision, (ii) lemma discovery

        • 6
        • No Conjecture (ii) (iii) (iv) T27 rev (X) = qrev(X, nil) G1

          T29 rev (qrev(X, nil)) = X G3,G4

          T30 rev (rev(X) <> nil) = X G5,G6

          T32 rotate (len(X), X) = X G9

          T33 f ac (X) = qf ac(X, 1) G10

          • 7

            T35 exp (X, Y ) = qexp(X, Y, 1) G12 T36

          • X ∈ Y → X ∈ (Y <> Z) th

            ⋆ X ∈ n ⋆ T39 (Y, Z) → X ∈ Z

            T40 X ⊂ Y → (X ∪ Y = Y ) ⋆ X ∈ Y → X ∈ ⋆ T42 (Y ∪ Z)

            T45 X ∈ insert (X, Y ) ⋆ len ⋆ T48 (isort(X)) = len(X) L18

            T49 X ∈ isort (Y ) → X ∈ Y L19 ⋆ count ⋆ T50 (X, isort(Y )) = count(X, Y ) L20, L21

            (ii) lemma discovery, (iii) generalization, (iv) casesplit No Lemma L1 X + s(Y ) = s(X + Y ) L3 len (X <> Y :: nil) = s(len(X))

            

          L8 rev (X <> (Y :: nil)) = Y :: rev(X)

          L10 rev ((X <> Y :: nil) <> nil) = Y :: rev(X <> nil) L11 (X <> (Y :: nil)) <> Z = X <> (Y :: Z)

          L12 ordered (Y ) → ordered(insert(X, Y ))

          L14 even (len(W <> Z)) → even(len(W <> X :: Y :: Z))

          L15 len (X <> Y :: Z :: nil) = s(s(len(X)))

          L17 half (X + s(s(Y ))) = s(half (X + Y ))

          L18 len (insert(X, Y )) = s(len(Y )) L19

            X

            6 = Y → (X ∈ insert(Y, Z) → X ∈ Z)

          L20 count (X, insert(X, Y )) = s(count(X, Y ))

          L21

            X

            6 = Y → (count(X, insert(Y, Z)) = count(X, Z))

          L22 (X <> Y ) <> Z = X <> (Y <> Z)

          L23 (X ∗ Y ) ∗ Z = X ∗ (Y ∗ Z) L24 (X + Y ) + Z = X + (Y + Z) 7

            Results: Generalizations No Generalization Lemmata G1 rev (X) <> Y = qrev(X, Y ) L22

            G2 revf lat (X) <> Y = qrevf lat(X, Y ) L22 G3 rev (qrev(X, Y )) = rev(Y ) <> X L11

            G4 rev (qrev(X, rev(Y ))) = Y <> X L8,L11 G5 rev (rev(X) <> Y ) = rev(Y ) <> X L11

            G6 rev (rev(X) <> rev(Y )) = Y <> X L8,L11 G7 qrev (qrev(X, Y ), nil) = rev(Y ) <> X L11 G8 qrev (qrev(X, rev(Y )), nil) = Y <> X L8,L11 G9 rotate (len(X), X <> Y ) = Y <> X L11, L22

            G10 f ac (X) ∗ Y = qf ac(X, Y ) L23 G11 (X ∗ Y ) + Z = mult(X, Y, Z) L24 G12 exp (X, Y ) ∗ Z = qexp(X, Y, Z) L23 7 special search problems .

          • Negative theoretical results create

            

          common in practice:

          • These problems induction rule choice, lemmas & generalisations.

            rippling .

          • Proof plan for induction based on
          • 7 • Rippling: selective; bidirectional; terminating and offers heuristic solution to special problems .

            Ripple analysis : Induction rules chosen to suit rippling. –

            • – Different patterns of proof breakdown suggest different patches .
            • – Ripple breakdowns suggest: induction revision; lemma speculation or generalisation .
            • – Implemented via proof planning with critics.

              Practical Session

            • Attempt to prove theorems with XBarnacle and IsaPlanner
            • 7 systems. • XBarnacle rather old but easier to use.

                Interactive gui on old Clam proof planner.

              • IsaPlanner latest experimental version, but functionality not yet complete.

                Proof planner built on Isabelle prover.