Theory Extraction

(*  Title:      HOL/Extraction.thy
    Author:     Stefan Berghofer, TU Muenchen
*)

section ‹Program extraction for HOL›

theory Extraction
imports Option
begin

subsection ‹Setup›

setup Extraction.add_types
      [("bool", ([], NONE))] #>
  Extraction.set_preprocessor (fn thy =>
    Proofterm.rewrite_proof_notypes
      ([], Rewrite_HOL_Proof.elim_cong :: Proof_Rewrite_Rules.rprocs true) o
    Proofterm.rewrite_proof thy
      (Rewrite_HOL_Proof.rews,
       Proof_Rewrite_Rules.rprocs true @ [Proof_Rewrite_Rules.expand_of_class thy]) o
    Proof_Rewrite_Rules.elim_vars (curry Const const_namedefault))

lemmas [extraction_expand] =
  meta_spec atomize_eq atomize_all atomize_imp atomize_conj
  allE rev_mp conjE Eq_TrueI Eq_FalseI eqTrueI eqTrueE eq_cong2
  notE' impE' impE iffE imp_cong simp_thms eq_True eq_False
  induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq
  induct_atomize induct_atomize' induct_rulify induct_rulify'
  induct_rulify_fallback induct_trueI
  True_implies_equals implies_True_equals TrueE
  False_implies_equals implies_False_swap

lemmas [extraction_expand_def] =
  HOL.induct_forall_def HOL.induct_implies_def HOL.induct_equal_def HOL.induct_conj_def
  HOL.induct_true_def HOL.induct_false_def

datatype (plugins only: code extraction) sumbool = Left | Right

subsection ‹Type of extracted program›

extract_type
  "typeof (Trueprop P) ≡ typeof P"

  "typeof P ≡ Type (TYPE(Null)) ⟹ typeof Q ≡ Type (TYPE('Q)) ⟹
     typeof (P ⟶ Q) ≡ Type (TYPE('Q))"

  "typeof Q ≡ Type (TYPE(Null)) ⟹ typeof (P ⟶ Q) ≡ Type (TYPE(Null))"

  "typeof P ≡ Type (TYPE('P)) ⟹ typeof Q ≡ Type (TYPE('Q)) ⟹
     typeof (P ⟶ Q) ≡ Type (TYPE('P ⇒ 'Q))"

  "(λx. typeof (P x)) ≡ (λx. Type (TYPE(Null))) ⟹
     typeof (∀x. P x) ≡ Type (TYPE(Null))"

  "(λx. typeof (P x)) ≡ (λx. Type (TYPE('P))) ⟹
     typeof (∀x::'a. P x) ≡ Type (TYPE('a ⇒ 'P))"

  "(λx. typeof (P x)) ≡ (λx. Type (TYPE(Null))) ⟹
     typeof (∃x::'a. P x) ≡ Type (TYPE('a))"

  "(λx. typeof (P x)) ≡ (λx. Type (TYPE('P))) ⟹
     typeof (∃x::'a. P x) ≡ Type (TYPE('a × 'P))"

  "typeof P ≡ Type (TYPE(Null)) ⟹ typeof Q ≡ Type (TYPE(Null)) ⟹
     typeof (P ∨ Q) ≡ Type (TYPE(sumbool))"

  "typeof P ≡ Type (TYPE(Null)) ⟹ typeof Q ≡ Type (TYPE('Q)) ⟹
     typeof (P ∨ Q) ≡ Type (TYPE('Q option))"

  "typeof P ≡ Type (TYPE('P)) ⟹ typeof Q ≡ Type (TYPE(Null)) ⟹
     typeof (P ∨ Q) ≡ Type (TYPE('P option))"

  "typeof P ≡ Type (TYPE('P)) ⟹ typeof Q ≡ Type (TYPE('Q)) ⟹
     typeof (P ∨ Q) ≡ Type (TYPE('P + 'Q))"

  "typeof P ≡ Type (TYPE(Null)) ⟹ typeof Q ≡ Type (TYPE('Q)) ⟹
     typeof (P ∧ Q) ≡ Type (TYPE('Q))"

  "typeof P ≡ Type (TYPE('P)) ⟹ typeof Q ≡ Type (TYPE(Null)) ⟹
     typeof (P ∧ Q) ≡ Type (TYPE('P))"

  "typeof P ≡ Type (TYPE('P)) ⟹ typeof Q ≡ Type (TYPE('Q)) ⟹
     typeof (P ∧ Q) ≡ Type (TYPE('P × 'Q))"

  "typeof (P = Q) ≡ typeof ((P ⟶ Q) ∧ (Q ⟶ P))"

  "typeof (x ∈ P) ≡ typeof P"

subsection ‹Realizability›

realizability
  "(realizes t (Trueprop P)) ≡ (Trueprop (realizes t P))"

  "(typeof P) ≡ (Type (TYPE(Null))) ⟹
     (realizes t (P ⟶ Q)) ≡ (realizes Null P ⟶ realizes t Q)"

  "(typeof P) ≡ (Type (TYPE('P))) ⟹
   (typeof Q) ≡ (Type (TYPE(Null))) ⟹
     (realizes t (P ⟶ Q)) ≡ (∀x::'P. realizes x P ⟶ realizes Null Q)"

  "(realizes t (P ⟶ Q)) ≡ (∀x. realizes x P ⟶ realizes (t x) Q)"

  "(λx. typeof (P x)) ≡ (λx. Type (TYPE(Null))) ⟹
     (realizes t (∀x. P x)) ≡ (∀x. realizes Null (P x))"

  "(realizes t (∀x. P x)) ≡ (∀x. realizes (t x) (P x))"

  "(λx. typeof (P x)) ≡ (λx. Type (TYPE(Null))) ⟹
     (realizes t (∃x. P x)) ≡ (realizes Null (P t))"

  "(realizes t (∃x. P x)) ≡ (realizes (snd t) (P (fst t)))"

  "(typeof P) ≡ (Type (TYPE(Null))) ⟹
   (typeof Q) ≡ (Type (TYPE(Null))) ⟹
     (realizes t (P ∨ Q)) ≡
     (case t of Left ⇒ realizes Null P | Right ⇒ realizes Null Q)"

  "(typeof P) ≡ (Type (TYPE(Null))) ⟹
     (realizes t (P ∨ Q)) ≡
     (case t of None ⇒ realizes Null P | Some q ⇒ realizes q Q)"

  "(typeof Q) ≡ (Type (TYPE(Null))) ⟹
     (realizes t (P ∨ Q)) ≡
     (case t of None ⇒ realizes Null Q | Some p ⇒ realizes p P)"

  "(realizes t (P ∨ Q)) ≡
   (case t of Inl p ⇒ realizes p P | Inr q ⇒ realizes q Q)"

  "(typeof P) ≡ (Type (TYPE(Null))) ⟹
     (realizes t (P ∧ Q)) ≡ (realizes Null P ∧ realizes t Q)"

  "(typeof Q) ≡ (Type (TYPE(Null))) ⟹
     (realizes t (P ∧ Q)) ≡ (realizes t P ∧ realizes Null Q)"

  "(realizes t (P ∧ Q)) ≡ (realizes (fst t) P ∧ realizes (snd t) Q)"

  "typeof P ≡ Type (TYPE(Null)) ⟹
     realizes t (¬ P) ≡ ¬ realizes Null P"

  "typeof P ≡ Type (TYPE('P)) ⟹
     realizes t (¬ P) ≡ (∀x::'P. ¬ realizes x P)"

  "typeof (P::bool) ≡ Type (TYPE(Null)) ⟹
   typeof Q ≡ Type (TYPE(Null)) ⟹
     realizes t (P = Q) ≡ realizes Null P = realizes Null Q"

  "(realizes t (P = Q)) ≡ (realizes t ((P ⟶ Q) ∧ (Q ⟶ P)))"

subsection ‹Computational content of basic inference rules›

theorem disjE_realizer:
  assumes r: "case x of Inl p  P p | Inr q  Q q"
  and r1: "p. P p  R (f p)" and r2: "q. Q q  R (g q)"
  shows "R (case x of Inl p  f p | Inr q  g q)"
proof (cases x)
  case Inl
  with r show ?thesis by simp (rule r1)
next
  case Inr
  with r show ?thesis by simp (rule r2)
qed

theorem disjE_realizer2:
  assumes r: "case x of None  P | Some q  Q q"
  and r1: "P  R f" and r2: "q. Q q  R (g q)"
  shows "R (case x of None  f | Some q  g q)"
proof (cases x)
  case None
  with r show ?thesis by simp (rule r1)
next
  case Some
  with r show ?thesis by simp (rule r2)
qed

theorem disjE_realizer3:
  assumes r: "case x of Left  P | Right  Q"
  and r1: "P  R f" and r2: "Q  R g"
  shows "R (case x of Left  f | Right  g)"
proof (cases x)
  case Left
  with r show ?thesis by simp (rule r1)
next
  case Right
  with r show ?thesis by simp (rule r2)
qed

theorem conjI_realizer:
  "P p  Q q  P (fst (p, q))  Q (snd (p, q))"
  by simp

theorem exI_realizer:
  "P y x  P (snd (x, y)) (fst (x, y))" by simp

theorem exE_realizer: "P (snd p) (fst p) 
  (x y. P y x  Q (f x y))  Q (let (x, y) = p in f x y)"
  by (cases p) (simp add: Let_def)

theorem exE_realizer': "P (snd p) (fst p) 
  (x y. P y x  Q)  Q" by (cases p) simp

realizers
  impI (P, Q): "λpq. pq"
    "λ(c: _) (d: _) P Q pq (h: _). allI ⋅ _ ∙ c ∙ (λx. impI ⋅ _ ⋅ _ ∙ (h ⋅ x))"

  impI (P): "Null"
    "λ(c: _) P Q (h: _). allI ⋅ _ ∙ c ∙ (λx. impI ⋅ _ ⋅ _ ∙ (h ⋅ x))"

  impI (Q): "λq. q" "λ(c: _) P Q q. impI ⋅ _ ⋅ _"

  impI: "Null" "impI"

  mp (P, Q): "λpq. pq"
    "λ(c: _) (d: _) P Q pq (h: _) p. mp ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ p ∙ c ∙ h)"

  mp (P): "Null"
    "λ(c: _) P Q (h: _) p. mp ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ p ∙ c ∙ h)"

  mp (Q): "λq. q" "λ(c: _) P Q q. mp ⋅ _ ⋅ _"

  mp: "Null" "mp"

  allI (P): "λp. p" "λ(c: _) P (d: _) p. allI ⋅ _ ∙ d"

  allI: "Null" "allI"

  spec (P): "λx p. p x" "λ(c: _) P x (d: _) p. spec ⋅ _ ⋅ x ∙ d"

  spec: "Null" "spec"

  exI (P): "λx p. (x, p)" "λ(c: _) P x (d: _) p. exI_realizer ⋅ P ⋅ p ⋅ x ∙ c ∙ d"

  exI: "λx. x" "λP x (c: _) (h: _). h"

  exE (P, Q): "λp pq. let (x, y) = p in pq x y"
    "λ(c: _) (d: _) P Q (e: _) p (h: _) pq. exE_realizer ⋅ P ⋅ p ⋅ Q ⋅ pq ∙ c ∙ e ∙ d ∙ h"

  exE (P): "Null"
    "λ(c: _) P Q (d: _) p. exE_realizer' ⋅ _ ⋅ _ ⋅ _ ∙ c ∙ d"

  exE (Q): "λx pq. pq x"
    "λ(c: _) P Q (d: _) x (h1: _) pq (h2: _). h2 ⋅ x ∙ h1"

  exE: "Null"
    "λP Q (c: _) x (h1: _) (h2: _). h2 ⋅ x ∙ h1"

  conjI (P, Q): "Pair"
    "λ(c: _) (d: _) P Q p (h: _) q. conjI_realizer ⋅ P ⋅ p ⋅ Q ⋅ q ∙ c ∙ d ∙ h"

  conjI (P): "λp. p"
    "λ(c: _) P Q p. conjI ⋅ _ ⋅ _"

  conjI (Q): "λq. q"
    "λ(c: _) P Q (h: _) q. conjI ⋅ _ ⋅ _ ∙ h"

  conjI: "Null" "conjI"

  conjunct1 (P, Q): "fst"
    "λ(c: _) (d: _) P Q pq. conjunct1 ⋅ _ ⋅ _"

  conjunct1 (P): "λp. p"
    "λ(c: _) P Q p. conjunct1 ⋅ _ ⋅ _"

  conjunct1 (Q): "Null"
    "λ(c: _) P Q q. conjunct1 ⋅ _ ⋅ _"

  conjunct1: "Null" "conjunct1"

  conjunct2 (P, Q): "snd"
    "λ(c: _) (d: _) P Q pq. conjunct2 ⋅ _ ⋅ _"

  conjunct2 (P): "Null"
    "λ(c: _) P Q p. conjunct2 ⋅ _ ⋅ _"

  conjunct2 (Q): "λp. p"
    "λ(c: _) P Q p. conjunct2 ⋅ _ ⋅ _"

  conjunct2: "Null" "conjunct2"

  disjI1 (P, Q): "Inl"
    "λ(c: _) (d: _) P Q p. iffD2 ⋅ _ ⋅ _ ∙ (sum.case_1 ⋅ P ⋅ _ ⋅ p ∙ arity_type_bool ∙ c ∙ d)"

  disjI1 (P): "Some"
    "λ(c: _) P Q p. iffD2 ⋅ _ ⋅ _ ∙ (option.case_2 ⋅ _ ⋅ P ⋅ p ∙ arity_type_bool ∙ c)"

  disjI1 (Q): "None"
    "λ(c: _) P Q. iffD2 ⋅ _ ⋅ _ ∙ (option.case_1 ⋅ _ ⋅ _ ∙ arity_type_bool ∙ c)"

  disjI1: "Left"
    "λP Q. iffD2 ⋅ _ ⋅ _ ∙ (sumbool.case_1 ⋅ _ ⋅ _ ∙ arity_type_bool)"

  disjI2 (P, Q): "Inr"
    "λ(d: _) (c: _) Q P q. iffD2 ⋅ _ ⋅ _ ∙ (sum.case_2 ⋅ _ ⋅ Q ⋅ q ∙ arity_type_bool ∙ c ∙ d)"

  disjI2 (P): "None"
    "λ(c: _) Q P. iffD2 ⋅ _ ⋅ _ ∙ (option.case_1 ⋅ _ ⋅ _ ∙ arity_type_bool ∙ c)"

  disjI2 (Q): "Some"
    "λ(c: _) Q P q. iffD2 ⋅ _ ⋅ _ ∙ (option.case_2 ⋅ _ ⋅ Q ⋅ q ∙ arity_type_bool ∙ c)"

  disjI2: "Right"
    "λQ P. iffD2 ⋅ _ ⋅ _ ∙ (sumbool.case_2 ⋅ _ ⋅ _ ∙ arity_type_bool)"

  disjE (P, Q, R): "λpq pr qr.
     (case pq of Inl p ⇒ pr p | Inr q ⇒ qr q)"
    "λ(c: _) (d: _) (e: _) P Q R pq (h1: _) pr (h2: _) qr.
       disjE_realizer ⋅ _ ⋅ _ ⋅ pq ⋅ R ⋅ pr ⋅ qr ∙ c ∙ d ∙ e ∙ h1 ∙ h2"

  disjE (Q, R): "λpq pr qr.
     (case pq of None ⇒ pr | Some q ⇒ qr q)"
    "λ(c: _) (d: _) P Q R pq (h1: _) pr (h2: _) qr.
       disjE_realizer2 ⋅ _ ⋅ _ ⋅ pq ⋅ R ⋅ pr ⋅ qr ∙ c ∙ d ∙ h1 ∙ h2"

  disjE (P, R): "λpq pr qr.
     (case pq of None ⇒ qr | Some p ⇒ pr p)"
    "λ(c: _) (d: _) P Q R pq (h1: _) pr (h2: _) qr (h3: _).
       disjE_realizer2 ⋅ _ ⋅ _ ⋅ pq ⋅ R ⋅ qr ⋅ pr ∙ c ∙ d ∙ h1 ∙ h3 ∙ h2"

  disjE (R): "λpq pr qr.
     (case pq of Left ⇒ pr | Right ⇒ qr)"
    "λ(c: _) P Q R pq (h1: _) pr (h2: _) qr.
       disjE_realizer3 ⋅ _ ⋅ _ ⋅ pq ⋅ R ⋅ pr ⋅ qr ∙ c ∙ h1 ∙ h2"

  disjE (P, Q): "Null"
    "λ(c: _) (d: _) P Q R pq. disjE_realizer ⋅ _ ⋅ _ ⋅ pq ⋅ (λx. R) ⋅ _ ⋅ _ ∙ c ∙ d ∙ arity_type_bool"

  disjE (Q): "Null"
    "λ(c: _) P Q R pq. disjE_realizer2 ⋅ _ ⋅ _ ⋅ pq ⋅ (λx. R) ⋅ _ ⋅ _ ∙ c ∙ arity_type_bool"

  disjE (P): "Null"
    "λ(c: _) P Q R pq (h1: _) (h2: _) (h3: _).
       disjE_realizer2 ⋅ _ ⋅ _ ⋅ pq ⋅ (λx. R) ⋅ _ ⋅ _ ∙ c ∙ arity_type_bool ∙ h1 ∙ h3 ∙ h2"

  disjE: "Null"
    "λP Q R pq. disjE_realizer3 ⋅ _ ⋅ _ ⋅ pq ⋅ (λx. R) ⋅ _ ⋅ _ ∙ arity_type_bool"

  FalseE (P): "default"
    "λ(c: _) P. FalseE ⋅ _"

  FalseE: "Null" "FalseE"

  notI (P): "Null"
    "λ(c: _) P (h: _). allI ⋅ _ ∙ c ∙ (λx. notI ⋅ _ ∙ (h ⋅ x))"

  notI: "Null" "notI"

  notE (P, R): "λp. default"
    "λ(c: _) (d: _) P R (h: _) p. notE ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ p ∙ c ∙ h)"

  notE (P): "Null"
    "λ(c: _) P R (h: _) p. notE ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ p ∙ c ∙ h)"

  notE (R): "default"
    "λ(c: _) P R. notE ⋅ _ ⋅ _"

  notE: "Null" "notE"

  subst (P): "λs t ps. ps"
    "λ(c: _) s t P (d: _) (h: _) ps. subst ⋅ s ⋅ t ⋅ P ps ∙ d ∙ h"

  subst: "Null" "subst"

  iffD1 (P, Q): "fst"
    "λ(d: _) (c: _) Q P pq (h: _) p.
       mp ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ p ∙ d ∙ (conjunct1 ⋅ _ ⋅ _ ∙ h))"

  iffD1 (P): "λp. p"
    "λ(c: _) Q P p (h: _). mp ⋅ _ ⋅ _ ∙ (conjunct1 ⋅ _ ⋅ _ ∙ h)"

  iffD1 (Q): "Null"
    "λ(c: _) Q P q1 (h: _) q2.
       mp ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ q2 ∙ c ∙ (conjunct1 ⋅ _ ⋅ _ ∙ h))"

  iffD1: "Null" "iffD1"

  iffD2 (P, Q): "snd"
    "λ(c: _) (d: _) P Q pq (h: _) q.
       mp ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ q ∙ d ∙ (conjunct2 ⋅ _ ⋅ _ ∙ h))"

  iffD2 (P): "λp. p"
    "λ(c: _) P Q p (h: _). mp ⋅ _ ⋅ _ ∙ (conjunct2 ⋅ _ ⋅ _ ∙ h)"

  iffD2 (Q): "Null"
    "λ(c: _) P Q q1 (h: _) q2.
       mp ⋅ _ ⋅ _ ∙ (spec ⋅ _ ⋅ q2 ∙ c ∙ (conjunct2 ⋅ _ ⋅ _ ∙ h))"

  iffD2: "Null" "iffD2"

  iffI (P, Q): "Pair"
    "λ(c: _) (d: _) P Q pq (h1 : _) qp (h2 : _). conjI_realizer ⋅
       (λpq. ∀x. P x ⟶ Q (pq x)) ⋅ pq ⋅
       (λqp. ∀x. Q x ⟶ P (qp x)) ⋅ qp ∙
       (arity_type_fun ∙ c ∙ d) ∙
       (arity_type_fun ∙ d ∙ c) ∙
       (allI ⋅ _ ∙ c ∙ (λx. impI ⋅ _ ⋅ _ ∙ (h1 ⋅ x))) ∙
       (allI ⋅ _ ∙ d ∙ (λx. impI ⋅ _ ⋅ _ ∙ (h2 ⋅ x)))"

  iffI (P): "λp. p"
    "λ(c: _) P Q (h1 : _) p (h2 : _). conjI ⋅ _ ⋅ _ ∙
       (allI ⋅ _ ∙ c ∙ (λx. impI ⋅ _ ⋅ _ ∙ (h1 ⋅ x))) ∙
       (impI ⋅ _ ⋅ _ ∙ h2)"

  iffI (Q): "λq. q"
    "λ(c: _) P Q q (h1 : _) (h2 : _). conjI ⋅ _ ⋅ _ ∙
       (impI ⋅ _ ⋅ _ ∙ h1) ∙
       (allI ⋅ _ ∙ c ∙ (λx. impI ⋅ _ ⋅ _ ∙ (h2 ⋅ x)))"

  iffI: "Null" "iffI"

end