Nominal/activities/tphols09/IDW/MW-Ex3.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 24 Sep 2018 12:24:38 +0100
changeset 559 d0438f49c683
parent 415 f1be8028a4a9
permissions -rw-r--r--
updated

theory Ex1
imports Main
begin

lemma True
proof

  subsect {* Basic context elements *}

  {
    fix x
    have "B x" sorry
  }
  thm this

  {
    assume A
    have B sorry
  }
  thm this

  {
    def x == a
    have "B x" sorry
  }
  thm this

  {
    obtain a where "B a" sorry
    have C sorry
  }
  thm this


  subsect {* Contexts vs. rules *}

  {
    fix A B
    fix x :: 'a
    assume "A x"
    have "B x" sorry
  }
  note r = this

  {
    ML_prf {* val ctxt = @{context} *}
    ML_prf {* val (((Ts, ts), [th']), ctxt') =
      Variable.import_thms true [@{thm r}] ctxt *}
  }

qed


section {* Building up contexts in ML *}

subsection {* Variables *}

lemma True
proof

  ML_prf {* val ctxt0 = @{context} *}
  ML_prf {* val (xs, ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y", "z"] *}
  ML_prf {* val t = Syntax.read_term ctxt1 "(x::nat) = y + z" *}
  ML_prf {* val ctxt2 = ctxt1 |> Variable.declare_term t *}
  ML_prf {* Syntax.read_term ctxt2 "x" *}

qed

ML {* val ctxt0 = @{context} *}
ML {* val (xs, ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y", "z"] *}
ML {* val (xs, ctxt1) = ctxt0 |> Variable.variant_fixes ["x", "y", "z"] *}

ML {* Variable.focus (Thm.cprem_of @{thm exE} 2) @{context} *}


subsection {* Assumptions *}

lemma True
proof

ML_prf {* val ctxt0 = @{context} *}
ML_prf {*
  val (asms, ctxt1) = ctxt0
    |> Assumption.add_assumes [@{cprop "x = y"}, @{cprop "y = z"}]
*}
ML_prf {* Assumption.export false ctxt1 ctxt0 (@{thm trans} OF asms) *}


ML_prf {* val ctxt0 = @{context} *}
ML_prf {* val (_, ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y", "z"] *}
ML_prf {* val ts = Syntax.read_props ctxt1 ["x = y", "y = z"] *}
ML_prf {* val ctxt2 = ctxt1 |> fold Variable.declare_term ts *}

ML_prf {*
  val (asms, ctxt3) =
    ctxt2
    |> Assumption.add_assumes (map (cterm_of (ProofContext.theory_of ctxt2)) ts)
*}
ML_prf {* Assumption.export false ctxt3 ctxt0 (@{thm trans} OF asms) *}
ML_prf {* ProofContext.export ctxt3 ctxt0 [@{thm trans} OF asms] *}

qed


subsection {* Local definitions (non-polymorphic) *}

lemma True
proof
  fix x y z :: nat

  ML_prf {* val ctxt0 = @{context} *}
  ML_prf {* val ((a, a_def), ctxt1) =
    ctxt0 |> LocalDefs.add_def ((@{binding a}, NoSyn), @{term "x + y"})
  *}
  ML_prf {* val ((b, b_def), ctxt2) =
    ctxt1 |> LocalDefs.add_def ((@{binding b}, NoSyn), @{term "y + z"})
  *}

  ML_prf {*
    Goal.prove ctxt2 [] [] (Syntax.read_prop ctxt2 "a + b = x + 2 * y + z")
      (fn _ => asm_full_simp_tac (local_simpset_of ctxt2 addsimps [a_def, b_def]) 1)
  *}

qed


subsection {* Local elimination *}

lemma True
proof
  assume ex: "EX x. B x"

  ML_prf {* val ctxt0 = @{context} *}
  ML_prf {* val (([x], [B]), ctxt1) = ctxt0
    |> Obtain.result (fn _ => etac @{thm exE} 1) @{thms ex} *}
  ML_prf {*
    ProofContext.export ctxt1 ctxt0 [Thm.reflexive x]
      handle ERROR msg => (warning msg; []) *}

qed

end