Tutorial/Tutorial1.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 20 Jun 2011 20:08:16 +0900
changeset 2878 06d91b7b5756
parent 2706 8ae1c2e6369e
child 3132 87eca760dcba
permissions -rw-r--r--
Abs_set_fcb
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
header {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
     3
  Nominal Isabelle Tutorial at POPL'11
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
     4
  ====================================
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  Nominal Isabelle is a definitional extension of Isabelle/HOL, which
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
     7
  means it does not add any new axioms to higher-order logic. It just
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  adds new definitions and an infrastructure for 'nominal resoning'.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
    11
  The jEdit Interface
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
    12
  -------------------
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    14
  The Isabelle theorem prover comes with an interface for jEdit. 
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  Unlike many other theorem prover interfaces (e.g. ProofGeneral) where you 
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    16
  advance a 'checked' region in a proof script, this interface immediately 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    17
  checks the whole buffer. The text you type is also immediately checked. 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    18
  Malformed text or unfinished proofs are highlighted in red with a little 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    19
  red 'stop' signal on the left-hand side. If you drag the 'red-box' cursor 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    20
  over a line, the Output window gives further feedback. 
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  Note: If a section is not parsed correctly, the work-around is to cut it 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  out and paste it back into the text (cut-out: Ctrl + X; paste in: Ctrl + V;
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  Cmd is Ctrl on the Mac)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    26
  Nominal Isabelle and jEdit can be started by typing on the command line
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
     isabelle jedit -l HOL-Nominal2
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
     isabelle jedit -l HOL-Nominal2 A.thy B.thy ...
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  Symbols
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  ------- 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  Symbols can considerably improve the readability of your statements and proofs. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  They can be input by just typing 'name-of-symbol' where 'name-of-symbol' is the 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  usual latex name of that symbol. A little window will then appear in which 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  you can select the symbol. With `Escape' you can ignore any suggestion.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  There are some handy short-cuts for frequently used symbols. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  For example 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
      short-cut  symbol
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
          =>      \<Rightarrow>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
          ==>     \<Longrightarrow>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
          -->     \<longrightarrow>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
          !       \<forall>        
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
          ?       \<exists>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
          /\      \<and>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
          \/      \<or>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
          ~       \<not>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
          ~=      \<noteq>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
          :       \<in>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
          ~:      \<notin>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    57
  For nominal the following two symbols have a special meaning
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    59
        \<sharp>    sharp     (freshness)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    60
        \<bullet>    bullet    (permutation application)
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
    63
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
theory Tutorial1
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
imports Lambda
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
begin
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
section {* Theories *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  All formal developments in Isabelle are part of a theory. A theory 
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
    72
  needs to have a name and must import some pre-existing theory. The 
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    73
  imported theory will normally be Nominal2 (which provides many useful 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    74
  concepts like set-theory, lists, nominal things etc). For the purpose 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    75
  of this tutorial we import the theory Lambda.thy, which contains 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
    76
  already some useful definitions for (alpha-equated) lambda terms.
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
section {* Types *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
text {*
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
    84
  Isabelle is based on simple types including some polymorphism. It also 
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
    85
  includes overloading, which means that sometimes explicit type annotations 
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
    86
  need to be given.
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
    88
    - Base types include: nat, bool, string, lam (defined in the Lambda theory)
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
    - Type formers include: 'a list, ('a \<times> 'b), 'c set   
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
    - Type variables are written like in ML with an apostrophe: 'a, 'b, \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  Types known to Isabelle can be queried using the command "typ".
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
typ nat
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
typ bool
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
typ string           
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   100
typ lam           -- {* alpha-equated lambda terms defined in Lambda.thy *}
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   101
typ name          -- {* type of (object) variables defined in Lambda.thy *}
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   102
typ "('a \<times> 'b)"   -- {* pair type *}
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   103
typ "'c set"      -- {* set type *}
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   104
typ "'a list"     -- {* list type *}
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   105
typ "lam \<Rightarrow> nat"   -- {* type of functions from lambda terms to natural numbers *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   108
text {* Some malformed types - note the "stop" signal on the left margin *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
(*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
typ boolean     -- {* undeclared type *} 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
typ set         -- {* type argument missing *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
*)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
section {* Terms *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
   Every term in Isabelle needs to be well-typed. However they can have 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
   polymorphic type. Whether a term is accepted can be queried using
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
   the command "term". 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
term c                 -- {* a variable of polymorphic type *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
term "1::nat"          -- {* the constant 1 in natural numbers *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
term 1                 -- {* the constant 1 with polymorphic type *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
term "{1, 2, 3::nat}"  -- {* the set containing natural numbers 1, 2 and 3 *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
term "[1, 2, 3]"       -- {* the list containing the polymorphic numbers 1, 2 and 3 *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
term "(True, ''c'')"   -- {* a pair consisting of the boolean true and the string "c" *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
term "Suc 0"           -- {* successor of 0, in other words 1::nat *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
term "Lam [x].Var x"   -- {* a lambda-term *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
term "App t1 t2"       -- {* another lambda-term *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
term "x::name"         -- {* an (object) variable of type name *}
2705
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2692
diff changeset
   134
term "atom (x::name)"  -- {* atom is an overloaded function *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
text {* 
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   137
  Lam [x].Var is the syntax we made up for lambda abstractions. If you
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   138
  prefer, you can have your own syntax (but \<lambda> is already taken up for 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   139
  Isabelle's functions). We also came up with the type "name" for variables. 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   140
  You can introduce your own types of object variables using the 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   141
  command atom_decl: 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   142
*}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   144
atom_decl ident
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   145
atom_decl ty_var
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   146
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   147
text {*
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   148
  Isabelle provides some useful colour feedback about its constants (black), 
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  free variables (blue) and bound variables (green).
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   152
term "True"    -- {* a constant defined somewhere...written in black *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
term "true"    -- {* not recognised as a constant, therefore it is interpreted
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
                     as a free variable, written in blue *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
term "\<forall>x. P x" -- {* x is bound (green), P is free (blue) *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
text {* Formulae in Isabelle/HOL are terms of type bool *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
term "True"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
term "True \<and> False"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
term "True \<or> B"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
term "{1,2,3} = {3,2,1}"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
term "\<forall>x. P x"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
term "A \<longrightarrow> B"
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   166
term "atom a \<sharp> t"   -- {* freshness in Nominal *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  When working with Isabelle, one deals with an object logic (that is HOL) and 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  Isabelle's rule framework (called Pure). Occasionally one has to pay attention 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  to this fact. But for the moment we ignore this completely.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
term "A \<longrightarrow> B"  -- {* versus *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
term "A \<Longrightarrow> B"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
term "\<forall>x. P x"  -- {* versus *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
term "\<And>x. P x"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   181
section {* Inductive Definitions: Evaluation Relation *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
2687
d0fb94035969 first split of tutorrial theory
Christian Urban <urbanc@in.tum.de>
parents: 2686
diff changeset
   183
text {*
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   184
  In this section we want to define inductively the evaluation
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   185
  relation and for cbv-reduction relation.
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   187
  Inductive definitions in Isabelle start with the keyword "inductive" 
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   188
  and a predicate name.  One can optionally provide a type for the predicate. 
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   189
  Clauses of the inductive predicate are introduced by "where" and more than 
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   190
  two clauses need to be separated by "|". One can also give a name to each 
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   191
  clause and indicate that it should be added to the hints database ("[intro]"). 
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   192
  A typical clause has some premises and a conclusion. This is written in 
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   193
  Isabelle as:
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
   "premise \<Longrightarrow> conclusion"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
   "premise1 \<Longrightarrow> premise2 \<Longrightarrow> \<dots> premisen \<Longrightarrow> conclusion"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  There is an alternative way of writing the latter clause as
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
   "\<lbrakk>premise1; premise2; \<dots> premisen\<rbrakk> \<Longrightarrow> conclusion"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  If no premise is present, then one just writes
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
   "conclusion"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  Below we give two definitions for the transitive closure
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
inductive
2704
b4bad45dbc0f Down as infixr
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2694
diff changeset
   210
  eval :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixr "\<Down>" 60)
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
where
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   212
  e_Lam[intro]:  "Lam [x].t \<Down> Lam [x].t"
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   213
| e_App[intro]:  "\<lbrakk>t1 \<Down> Lam [x].t; t2 \<Down> v'; t[x::=v'] \<Down> v\<rbrakk> \<Longrightarrow> App t1 t2 \<Down> v"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   215
text {* 
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   216
  Values and cbv are also defined using inductive.  
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   217
*}
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   218
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
inductive
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   220
  val :: "lam \<Rightarrow> bool" 
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
where
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   222
  v_Lam[intro]:   "val (Lam [x].t)"
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   223
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
section {* Theorems *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  A central concept in Isabelle is that of theorems. Isabelle's theorem
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  database can be queried using 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   231
thm e_App
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   232
thm e_Lam
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
thm conjI
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
thm conjunct1
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
  Notice that theorems usually contain schematic variables (e.g. ?P, ?Q, \<dots>). 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  These schematic variables can be substituted with any term (of the right type 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  of course). 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
  When defining the predicates beta_star, Isabelle provides us automatically 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  with the following theorems that state how they can be introduced and what 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
  constitutes an induction over them. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   246
thm eval.intros
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   247
thm eval.induct
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
section {* Lemmas / Theorems / Corollaries *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  Whether to use lemma, theorem or corollary makes no semantic difference 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  in Isabelle. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  A lemma starts with "lemma" and consists of a statement ("shows \<dots>") and 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  optionally a lemma name, some type-information for variables ("fixes \<dots>") 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  and some assumptions ("assumes \<dots>"). 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  Lemmas also need to have a proof, but ignore this 'detail' for the moment.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  Examples are
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
lemma alpha_equ:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  shows "Lam [x].Var x = Lam [y].Var y"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  by (simp add: lam.eq_iff Abs1_eq_iff lam.fresh fresh_at_base)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
lemma Lam_freshness:
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   270
  assumes a: "atom y \<sharp> Lam [x].t"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   271
  shows "(y = x) \<or> (y \<noteq> x \<and> atom y \<sharp> t)"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   272
  using a by (auto simp add: lam.fresh Abs_fresh_iff) 
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
lemma neutral_element:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  fixes x::"nat"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
  shows "x + 0 = x"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
  by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
  Note that in the last statement, the explicit type annotation is important 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  in order for Isabelle to be able to figure out what 0 stands for (e.g. a 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  natural number, a vector, etc) and which lemmas to apply.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
section {* Isar Proofs *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  Isar is a language for writing formal proofs that can be understood by humans 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
  and by Isabelle. An Isar proof can be thought of as a sequence of 'stepping stones' 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  that start with some assumptions and lead to the goal to be established. Every such 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  stepping stone is introduced by "have" followed by the statement of the stepping 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
  stone. An exception is the goal to be proved, which need to be introduced with "show".
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
      have "statement"  \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
      show "goal_to_be_proved" \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
  Since proofs usually do not proceed in a linear fashion, labels can be given 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
  to every stepping stone and they can be used later to explicitly refer to this 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
  corresponding stepping stone ("using").
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
      have label: "statement1"  \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
      \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
      have "later_statement" using label \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
  Each stepping stone (or have-statement) needs to have a justification. The 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  simplest justification is "sorry" which admits any stepping stone, even false 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  ones (this is good during the development of proofs). 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
      have "outrageous_false_statement" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
  Assumptions can be 'justified' using "by fact". 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
      have "assumption" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
  Derived facts can be justified using 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
      have "statement" by simp    -- simplification 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
      have "statement" by auto    -- proof search and simplification 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
      have "statement" by blast   -- only proof search 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
  If facts or lemmas are needed in order to justify a have-statement, then
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
  one can feed these facts into the proof by using "using label \<dots>" or 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
  "using theorem-name \<dots>". More than one label at a time is allowed. For
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
  example
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
      have "statement" using label1 label2 theorem_name by auto
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
  Induction proofs in Isar are set up by indicating over which predicate(s) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  the induction proceeds ("using a b") followed by the command "proof (induct)". 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
  In this way, Isabelle uses defaults for which induction should be performed. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  These defaults can be overridden by giving more information, like the variable 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
  over which a structural induction should proceed, or a specific induction principle, 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
  such as well-founded inductions. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
  After the induction is set up, the proof proceeds by cases. In Isar these 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  cases can be given in any order. Most commonly they are started with "case" and the 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
  name of the case, and optionally some legible names for the variables used inside 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
  the case. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
  In each "case", we need to establish a statement introduced by "show". Once 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
  this has been done, the next case can be started using "next". When all cases 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
  are completed, the proof can be finished using "qed".
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
  This means a typical induction proof has the following pattern
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
     proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
       case \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
       \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
       show \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
     next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
       case \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
       \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
       show \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
     \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
     qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   358
   Two very simple example proofs are as follows.
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   361
subsection {* EXERCISE 1 *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   363
lemma eval_val:
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   364
  assumes a: "val t"
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   365
  shows "t \<Down> t"
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   366
using a
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
proof (induct)
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   368
  case (v_Lam x t)
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   369
  show "Lam [x]. t \<Down> Lam [x]. t" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   372
subsection {* EXERCISE 2 *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   373
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   374
text {* Fill the gaps in the proof below. *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   375
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   376
lemma eval_to_val:
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   377
  assumes a: "t \<Down> t'"
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   378
  shows "val t'"
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   379
using a
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
proof (induct)
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   381
  case (e_Lam x t)
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   382
  show "val (Lam [x].t)" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
next
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   384
  case (e_App t1 x t t2 v v')
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   385
  -- {* all assumptions in this case *}
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   386
  have "t1 \<Down> Lam [x].t" by fact
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   387
  have ih1: "val (Lam [x]. t)" by fact
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   388
  have "t2 \<Down> v" by fact
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   389
  have ih2: "val v" by fact
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   390
  have "t [x ::= v] \<Down> v'" by fact
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   391
  have ih3: "val v'" by fact
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   392
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   393
  show "val v'" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
qed
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   395
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   396
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   397
section {* Datatypes: Evaluation Contexts*}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   398
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   399
text {*
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   400
  Datatypes can be defined in Isabelle as follows: we have to provide the name 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   401
  of the datatype and a list its type-constructors. Each type-constructor needs 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   402
  to have the information about the types of its arguments, and optionally 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   403
  can also contain some information about pretty syntax. For example, we like
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   404
  to write "\<box>" for holes.
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   405
*}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
  
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   407
datatype ctx = 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   408
    Hole ("\<box>")  
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   409
  | CAppL "ctx" "lam"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   410
  | CAppR "lam" "ctx"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   411
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   412
text {* Now Isabelle knows about: *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   413
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   414
typ ctx
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   415
term "\<box>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   416
term "CAppL"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   417
term "CAppL \<box> (Var x)"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   418
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   419
subsection {* MINI EXERCISE *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   420
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   421
text {*
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   422
  Try and see what happens if you apply a Hole to a Hole? 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   423
*}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   424
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   425
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   426
section {* Machines for Evaluation *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   427
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   428
type_synonym ctxs = "ctx list"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   429
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   430
inductive
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   431
  machine :: "lam \<Rightarrow> ctxs \<Rightarrow>lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto> <_,_>" [60, 60, 60, 60] 60)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   432
where
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   433
  m1: "<App t1 t2, Es> \<mapsto> <t1, (CAppL \<box> t2) # Es>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   434
| m2: "val v \<Longrightarrow> <v, (CAppL \<box> t2) # Es> \<mapsto> <t2, (CAppR v \<box>) # Es>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   435
| m3: "val v \<Longrightarrow> <v, (CAppR (Lam [x].t) \<box>) # Es> \<mapsto> <t[x ::= v], Es>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   436
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   437
text {*
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   438
  Since the machine defined above only performs a single reduction,
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   439
  we need to define the transitive closure of this machine. *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   440
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   441
inductive 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   442
  machines :: "lam \<Rightarrow> ctxs \<Rightarrow> lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto>* <_,_>" [60, 60, 60, 60] 60)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   443
where
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   444
  ms1: "<t,Es> \<mapsto>* <t,Es>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   445
| ms2: "\<lbrakk><t1,Es1> \<mapsto> <t2,Es2>; <t2,Es2> \<mapsto>* <t3,Es3>\<rbrakk> \<Longrightarrow> <t1,Es1> \<mapsto>* <t3,Es3>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   446
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   447
declare machine.intros[intro] machines.intros[intro]
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   448
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   449
section {* EXERCISE 3 *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   450
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   451
text {*
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   452
  We need to show that the machines-relation is transitive.
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   453
  Fill in the gaps below.   
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   454
*}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   455
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   456
lemma ms3[intro]:
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   457
  assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>" 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   458
  and     b: "<e2, Es2> \<mapsto>* <e3, Es3>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   459
  shows "<e1, Es1> \<mapsto>* <e3, Es3>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   460
using a b
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   461
proof(induct)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   462
  case (ms1 e1 Es1)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   463
  have c: "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   464
  
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   465
  show "<e1, Es1> \<mapsto>* <e3, Es3>" sorry
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   466
next
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   467
  case (ms2 e1 Es1 e2 Es2 e2' Es2') 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   468
  have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   469
  have d1: "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   470
  have d2: "<e1, Es1> \<mapsto> <e2, Es2>" by fact
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   471
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   472
  show "<e1, Es1> \<mapsto>* <e3, Es3>" sorry
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   473
qed
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   474
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
  Just like gotos in the Basic programming language, labels often reduce 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
  the readability of proofs. Therefore one can use in Isar the notation
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  "then have" in order to feed a have-statement to the proof of 
2705
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2692
diff changeset
   479
  the next have-statement. This is used in the second case below.
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
*}
2692
da9bed7baf23 better flow of proofs and definitions and proof
Christian Urban <urbanc@in.tum.de>
parents: 2690
diff changeset
   481
 
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   482
lemma 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   483
  assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>" 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   484
  and     b: "<e2, Es2> \<mapsto>* <e3, Es3>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   485
  shows "<e1, Es1> \<mapsto>* <e3, Es3>"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   486
using a b
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   487
proof(induct)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   488
  case (ms1 e1 Es1)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   489
  show "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   490
next
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   491
  case (ms2 e1 Es1 e2 Es2 e2' Es2')
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   492
  have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   493
  have "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   494
  then have d3: "<e2, Es2> \<mapsto>* <e3, Es3>" using ih by simp
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   495
  have d2: "<e1, Es1> \<mapsto> <e2, Es2>" by fact
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   496
  show "<e1, Es1> \<mapsto>* <e3, Es3>" using d2 d3 by auto
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   497
qed
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
  The label ih2 cannot be got rid of in this way, because it is used 
2705
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2692
diff changeset
   501
  two lines below and we cannot rearrange them. We can still avoid the
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
  label by feeding a sequence of facts into a proof using the 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
  "moreover"-chaining mechanism:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
   have "statement_1" \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
   moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
   have "statement_2" \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
   \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
   moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
   have "statement_n" \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
   ultimately have "statement" \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
  In this chain, all "statement_i" can be used in the proof of the final 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
  "statement". With this we can simplify our proof further to:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
  assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
  and     b: "<e2, Es2> \<mapsto>* <e3, Es3>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
  shows "<e1, Es1> \<mapsto>* <e3, Es3>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
using a b
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
proof(induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
  case (ms1 e1 Es1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
  show "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
  case (ms2 e1 Es1 e2 Es2 e2' Es2')
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
  have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
  have "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
  then have "<e2, Es2> \<mapsto>* <e3, Es3>" using ih by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
  moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
  have "<e1, Es1> \<mapsto> <e2, Es2>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
  ultimately show "<e1, Es1> \<mapsto>* <e3, Es3>" by auto
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   536
text {* 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   537
  While automatic proof procedures in Isabelle are not able to prove statements
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   538
  like "P = NP" assuming usual definitions for P and NP, they can automatically
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   539
  discharge the lemmas we just proved. For this we only have to set up the induction
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   540
  and auto will take care of the rest. This means we can write:
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   541
*}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   542
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   543
lemma 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   544
  assumes a: "val t"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   545
  shows "t \<Down> t"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   546
using a by (induct) (auto)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   547
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   548
lemma 
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   549
  assumes a: "t \<Down> t'"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   550
  shows "val t'"
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   551
using a by (induct) (auto)
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   552
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   553
lemma
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
  assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
  and     b: "<e2, Es2> \<mapsto>* <e3, Es3>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
  shows "<e1, Es1> \<mapsto>* <e3, Es3>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
using a b by (induct) (auto)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   559
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   560
section {* EXERCISE 4 *}
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   561
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   562
text {*
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   563
  The point of the machine is that it simulates the evaluation
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   564
  relation. Therefore we like to prove the following:
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   565
*}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
theorem 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
  assumes a: "t \<Down> t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
  shows "<t, []> \<mapsto>* <t', []>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
using a 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
proof (induct)
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   572
  case (e_Lam x t) 
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   573
  -- {* no assumptions *}
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents: 2688
diff changeset
   574
  show "<Lam [x].t, []> \<mapsto>* <Lam [x].t, []>" by auto
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
next
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   576
  case (e_App t1 x t t2 v' v) 
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   577
  -- {* all assumptions in this case *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
  have a1: "t1 \<Down> Lam [x].t" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
  have ih1: "<t1, []> \<mapsto>* <Lam [x].t, []>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
  have a2: "t2 \<Down> v'" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
  have ih2: "<t2, []> \<mapsto>* <v', []>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
  have a3: "t[x::=v'] \<Down> v" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
  have ih3: "<t[x::=v'], []> \<mapsto>* <v, []>" by fact
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   584
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   585
  -- {* your reasoning *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
  show "<App t1 t2, []> \<mapsto>* <v, []>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   589
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
  Again the automatic tools in Isabelle can discharge automatically 
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   592
  of the routine work in these proofs. We can write: 
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   593
*}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
theorem eval_implies_machines_ctx:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
  assumes a: "t \<Down> t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
  shows "<t, Es> \<mapsto>* <t', Es>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
by (induct arbitrary: Es)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
   (metis eval_to_val machine.intros ms1 ms2 ms3 v_Lam)+
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
corollary eval_implies_machines:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
  assumes a: "t \<Down> t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  shows "<t, []> \<mapsto>* <t', []>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
using a eval_implies_machines_ctx by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   608
section {* Function Definitions: Filling a Lambda-Term into a Context *}
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   609
 
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   610
text {*
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
  Many functions over datatypes can be defined by recursion on the
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
  structure. For this purpose, Isabelle provides "fun". To use it one needs 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
  to give a name for the function, its type, optionally some pretty-syntax 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
  and then some equations defining the function. Like in "inductive",
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
  "fun" expects that more than one such equation is separated by "|".
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
fun
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
  filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100, 100] 100)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
  "\<box>\<lbrakk>t\<rbrakk> = t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
| "(CAppL E t')\<lbrakk>t\<rbrakk> = App (E\<lbrakk>t\<rbrakk>) t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
| "(CAppR t' E)\<lbrakk>t\<rbrakk> = App t' (E\<lbrakk>t\<rbrakk>)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
2693
2abc8cb46a5c better version of Tutorial 1
Christian Urban <urbanc@in.tum.de>
parents: 2692
diff changeset
   625
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
  After this definition Isabelle will be able to simplify
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   628
  statements like: 
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   629
*}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   632
  shows "(CAppL \<box> (Var x))\<lbrakk>Var y\<rbrakk> = App (Var y) (Var x)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
  by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   635
fun 
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   636
 ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" (infixr "\<odot>" 99)
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
where
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   638
  "\<box> \<odot> E' = E'"
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   639
| "(CAppL E t') \<odot> E' = CAppL (E \<odot> E') t'"
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   640
| "(CAppR t' E) \<odot> E' = CAppR t' (E \<odot> E')"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
fun
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
  ctx_composes :: "ctxs \<Rightarrow> ctx" ("_\<down>" [110] 110)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
    "[]\<down> = \<box>"
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   646
  | "(E # Es)\<down> = (Es\<down>) \<odot> E"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   647
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
text {*  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
  Notice that we not just have given a pretty syntax for the functions, but
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   650
  also some precedences. The numbers inside the [\<dots>] stand for the precedences
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
  of the arguments; the one next to it the precedence of the whole term.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
  
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   653
  This means we have to write (Es1 \<odot> Es2) \<odot> Es3 otherwise Es1 \<odot> Es2 \<odot> Es3 is
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   654
  interpreted as Es1 \<odot> (Es2 \<odot> Es3).
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   655
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   657
section {* Structural Inductions over Contexts *}
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   658
 
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   659
text {*
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
  So far we have had a look at an induction over an inductive predicate. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
  Another important induction principle is structural inductions for 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
  datatypes. To illustrate structural inductions we prove some facts
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
  about context composition, some of which we will need later on.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   666
subsection {* EXERCISE 5 *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   668
text {* Complete the proof and remove the sorries. *}
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   670
lemma ctx_compose:
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   671
  shows "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
proof (induct E1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
  case Hole
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   674
  show "(\<box> \<odot> E2)\<lbrakk>t\<rbrakk> = \<box>\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
  case (CAppL E1 t')
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   677
  have ih: "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   678
  show "((CAppL E1 t') \<odot> E2)\<lbrakk>t\<rbrakk> = (CAppL E1 t')\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
  case (CAppR t' E1)
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   681
  have ih: "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   682
  show "((CAppR t' E1) \<odot> E2)\<lbrakk>t\<rbrakk> = (CAppR t' E1)\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   685
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   686
subsection {* EXERCISE 6 *}
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   687
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   688
text {*
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   689
  Remove the sorries in the ctx_append proof below. You can make 
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   690
  use of the following two properties.    
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   691
*}
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   692
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
lemma neut_hole:
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   694
  shows "E \<odot> \<box> = E"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
by (induct E) (simp_all)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   697
lemma compose_assoc:  
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   698
  shows "(E1 \<odot> E2) \<odot> E3 = E1 \<odot> (E2 \<odot> E3)"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
by (induct E1) (simp_all)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
lemma
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   702
  shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<odot> (Es1\<down>)"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
proof (induct Es1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
  case Nil
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   705
  show "([] @ Es2)\<down> = Es2\<down> \<odot> []\<down>" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
  case (Cons E Es1)
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   708
  have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<odot> Es1\<down>" by fact
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   710
  show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" sorry
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
  The last proof involves several steps of equational reasoning.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
  Isar provides some convenient means to express such equational 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
  reasoning in a much cleaner fashion using the "also have" 
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   717
  construction. 
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   718
*}
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   719
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   720
lemma 
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   721
  shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<odot> (Es1\<down>)"
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
proof (induct Es1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
  case Nil
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   724
  show "([] @ Es2)\<down> = Es2\<down> \<odot> []\<down>" using neut_hole by simp
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
  case (Cons E Es1)
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   727
  have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<odot> Es1\<down>" by fact
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   728
  have "((E # Es1) @ Es2)\<down> = (E # (Es1 @ Es2))\<down>" by simp
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   729
  also have "\<dots> = (Es1 @ Es2)\<down> \<odot> E" by simp
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   730
  also have "\<dots> = (Es2\<down> \<odot> Es1\<down>) \<odot> E" using ih by simp
2694
3485111c7d62 cleaned up tutorial1...added solution file
Christian Urban <urbanc@in.tum.de>
parents: 2693
diff changeset
   731
  also have "\<dots> = Es2\<down> \<odot> (Es1\<down> \<odot> E)" using compose_assoc by simp
2688
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   732
  also have "\<dots> = Es2\<down> \<odot> (E # Es1)\<down>" by simp
87b735f86060 a bit tuning
Christian Urban <urbanc@in.tum.de>
parents: 2687
diff changeset
   733
  finally show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" by simp
2686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   735
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   736
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
end  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738