Nominal/activities/tphols09/IDW/MW-Ex3.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 theory Ex1
       
     2 imports Main
       
     3 begin
       
     4 
       
     5 lemma True
       
     6 proof
       
     7 
       
     8   subsect {* Basic context elements *}
       
     9 
       
    10   {
       
    11     fix x
       
    12     have "B x" sorry
       
    13   }
       
    14   thm this
       
    15 
       
    16   {
       
    17     assume A
       
    18     have B sorry
       
    19   }
       
    20   thm this
       
    21 
       
    22   {
       
    23     def x == a
       
    24     have "B x" sorry
       
    25   }
       
    26   thm this
       
    27 
       
    28   {
       
    29     obtain a where "B a" sorry
       
    30     have C sorry
       
    31   }
       
    32   thm this
       
    33 
       
    34 
       
    35   subsect {* Contexts vs. rules *}
       
    36 
       
    37   {
       
    38     fix A B
       
    39     fix x :: 'a
       
    40     assume "A x"
       
    41     have "B x" sorry
       
    42   }
       
    43   note r = this
       
    44 
       
    45   {
       
    46     ML_prf {* val ctxt = @{context} *}
       
    47     ML_prf {* val (((Ts, ts), [th']), ctxt') =
       
    48       Variable.import_thms true [@{thm r}] ctxt *}
       
    49   }
       
    50 
       
    51 qed
       
    52 
       
    53 
       
    54 section {* Building up contexts in ML *}
       
    55 
       
    56 subsection {* Variables *}
       
    57 
       
    58 lemma True
       
    59 proof
       
    60 
       
    61   ML_prf {* val ctxt0 = @{context} *}
       
    62   ML_prf {* val (xs, ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y", "z"] *}
       
    63   ML_prf {* val t = Syntax.read_term ctxt1 "(x::nat) = y + z" *}
       
    64   ML_prf {* val ctxt2 = ctxt1 |> Variable.declare_term t *}
       
    65   ML_prf {* Syntax.read_term ctxt2 "x" *}
       
    66 
       
    67 qed
       
    68 
       
    69 ML {* val ctxt0 = @{context} *}
       
    70 ML {* val (xs, ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y", "z"] *}
       
    71 ML {* val (xs, ctxt1) = ctxt0 |> Variable.variant_fixes ["x", "y", "z"] *}
       
    72 
       
    73 ML {* Variable.focus (Thm.cprem_of @{thm exE} 2) @{context} *}
       
    74 
       
    75 
       
    76 subsection {* Assumptions *}
       
    77 
       
    78 lemma True
       
    79 proof
       
    80 
       
    81 ML_prf {* val ctxt0 = @{context} *}
       
    82 ML_prf {*
       
    83   val (asms, ctxt1) = ctxt0
       
    84     |> Assumption.add_assumes [@{cprop "x = y"}, @{cprop "y = z"}]
       
    85 *}
       
    86 ML_prf {* Assumption.export false ctxt1 ctxt0 (@{thm trans} OF asms) *}
       
    87 
       
    88 
       
    89 ML_prf {* val ctxt0 = @{context} *}
       
    90 ML_prf {* val (_, ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y", "z"] *}
       
    91 ML_prf {* val ts = Syntax.read_props ctxt1 ["x = y", "y = z"] *}
       
    92 ML_prf {* val ctxt2 = ctxt1 |> fold Variable.declare_term ts *}
       
    93 
       
    94 ML_prf {*
       
    95   val (asms, ctxt3) =
       
    96     ctxt2
       
    97     |> Assumption.add_assumes (map (cterm_of (ProofContext.theory_of ctxt2)) ts)
       
    98 *}
       
    99 ML_prf {* Assumption.export false ctxt3 ctxt0 (@{thm trans} OF asms) *}
       
   100 ML_prf {* ProofContext.export ctxt3 ctxt0 [@{thm trans} OF asms] *}
       
   101 
       
   102 qed
       
   103 
       
   104 
       
   105 subsection {* Local definitions (non-polymorphic) *}
       
   106 
       
   107 lemma True
       
   108 proof
       
   109   fix x y z :: nat
       
   110 
       
   111   ML_prf {* val ctxt0 = @{context} *}
       
   112   ML_prf {* val ((a, a_def), ctxt1) =
       
   113     ctxt0 |> LocalDefs.add_def ((@{binding a}, NoSyn), @{term "x + y"})
       
   114   *}
       
   115   ML_prf {* val ((b, b_def), ctxt2) =
       
   116     ctxt1 |> LocalDefs.add_def ((@{binding b}, NoSyn), @{term "y + z"})
       
   117   *}
       
   118 
       
   119   ML_prf {*
       
   120     Goal.prove ctxt2 [] [] (Syntax.read_prop ctxt2 "a + b = x + 2 * y + z")
       
   121       (fn _ => asm_full_simp_tac (local_simpset_of ctxt2 addsimps [a_def, b_def]) 1)
       
   122   *}
       
   123 
       
   124 qed
       
   125 
       
   126 
       
   127 subsection {* Local elimination *}
       
   128 
       
   129 lemma True
       
   130 proof
       
   131   assume ex: "EX x. B x"
       
   132 
       
   133   ML_prf {* val ctxt0 = @{context} *}
       
   134   ML_prf {* val (([x], [B]), ctxt1) = ctxt0
       
   135     |> Obtain.result (fn _ => etac @{thm exE} 1) @{thms ex} *}
       
   136   ML_prf {*
       
   137     ProofContext.export ctxt1 ctxt0 [Thm.reflexive x]
       
   138       handle ERROR msg => (warning msg; []) *}
       
   139 
       
   140 qed
       
   141 
       
   142 end