Tutorial/Tutorial1.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 19 Jan 2011 23:58:12 +0100
changeset 2686 52e1e98edb34
child 2687 d0fb94035969
permissions -rw-r--r--
added a very rough version of the tutorial; all seems to work
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
header {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
  Nominal Isabelle Tutorial
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  There will be hands-on exercises throughout the tutorial. Therefore
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  it would be good if you have your own laptop.
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
  Nominal Isabelle is a definitional extension of Isabelle/HOL, which
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  means it does not add any new axioms to higher-order logic. It merely
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  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
    13
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  The Interface
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  -------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  The Isabelle theorem prover comes with an interface for jEdit interface. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  Unlike many other theorem prover interfaces (e.g. ProofGeneral) where you 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  try to advance a 'checked' region in a proof script, this interface immediately 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  checks the whole buffer. The text you type is also immediately checked
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  as you type. Malformed text or unfinished proofs are highlighted in red 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  with a little red 'stop' signal on the left-hand side. If you drag the 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  'red-box' cursor over a line, the Output window gives further feedback. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  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
    27
  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
    28
  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
    29
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  Nominal Isabelle and the interface can be started on the command line with
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
     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
    33
     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
    34
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  Symbols
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  ------- 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  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
    40
  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
    41
  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
    42
  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
    43
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  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
    45
  For example 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      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
    48
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
          =>      \<Rightarrow>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
          ==>     \<Longrightarrow>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
          -->     \<longrightarrow>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
          !       \<forall>        
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
          ?       \<exists>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
          /\      \<and>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
          \/      \<or>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
          ~       \<not>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
          ~=      \<noteq>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
          :       \<in>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
          ~:      \<notin>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  For nominal two important symbols are
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
          \<sharp>       sharp     (freshness)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
          \<bullet>       bullet    (permutations)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
theory Tutorial1
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
imports Lambda
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
begin
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
section {* Theories *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  All formal developments in Isabelle are part of a theory. A theory 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  needs to have a name and must import some pre-existing theory (as indicated 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  above). The imported theory will normally be the theory Nominal2 (which 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  contains many useful concepts like set-theory, lists, nominal theory etc).
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
  For the purpose of the tutorial we import the theory Lambda.thy which
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  contains already some useful definitions for (alpha-equated) lambda terms.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
*}
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
section {* Types *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  Isabelle is based on simple types including some polymorphism. It also includes 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
  some overloading, which means that sometimes explicit type annotations need to 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
  be given.
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
    - Base types include: nat, bool, string
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
    - 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
    95
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
    - 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
    97
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
  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
    99
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
typ nat
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
typ bool
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
typ string           
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
typ lam             -- {* alpha-equated lambda terms defined in Lambda.thy *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
typ name            -- {* type of (object) variables defined in Lambda.thy *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
typ "('a \<times> 'b)"     -- {* pair type *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
typ "'c set"        -- {* set type *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
typ "'a list"       -- {* list type *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
typ "nat \<Rightarrow> bool"   -- {* type of functions from natural numbers to booleans *}
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
text {* Some malformed types: *}
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
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
   116
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
   117
*)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
section {* Terms *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
   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
   124
   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
   125
   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
   126
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
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
   129
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
   130
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
   131
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
   132
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
   133
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
   134
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
   135
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
   136
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
   137
term "x::name"         -- {* an (object) variable of type name *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
term "atom (x::name)"  -- {* atom is an overloded function *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  Lam [x].Var is the syntax we made up for lambda abstractions. You can have
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
  your own syntax. We also came up with "name". If you prefer, you can call
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  it identifiers or have more than one type for (object languag) variables.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  Isabelle provides some useful colour feedback about constants (black), 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  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
   147
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
term "True"    -- {* a constant that is defined in HOL...written in black *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
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
   151
                     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
   152
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
   153
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
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
   156
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
term "True"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
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
   159
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
   160
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
   161
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
   162
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
   163
term "atom a \<sharp> t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  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
   167
  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
   168
  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
   169
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
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
   172
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
   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 "\<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
   175
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
   176
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
section {* Inductive Definitions: Transitive Closures of Beta-Reduction *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
  The theory Lmabda alraedy contains a definition for beta-reduction, written
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
     t \<longrightarrow>b t'
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  In this section we want to define inductively the transitive closure of
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
  beta-reduction.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
  Inductive definitions in Isabelle start with the keyword "inductive" and a predicate 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  name.  One can optionally provide a type for the predicate. Clauses of the inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  predicate are introduced by "where" and more than two clauses need to be 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
  separated by "|". One can also give a name to each clause and indicate that it 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  should be added to the hints database ("[intro]"). A typical clause has some 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  premises and a conclusion. This is written in Isabelle as:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
   "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
   197
   "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
   198
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  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
   200
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
   "\<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
   202
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
  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
   204
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
   "conclusion"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  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
   208
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  beta_star1 :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<longrightarrow>b* _" [60, 60] 60)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
  bs1_refl: "t \<longrightarrow>b* t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
| bs1_trans: "\<lbrakk>t1 \<longrightarrow>b t2; t2 \<longrightarrow>b* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  beta_star2 :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<longrightarrow>b** _" [60, 60] 60)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  bs2_refl: "t \<longrightarrow>b** t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
| bs2_step: "t1 \<longrightarrow>b t2 \<Longrightarrow> t1 \<longrightarrow>b** t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
| bs2_trans: "\<lbrakk>t1 \<longrightarrow>b** t2; t2 \<longrightarrow>b** t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>b** t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
thm bs1_refl
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
thm bs2_trans
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
thm beta_star1.intros
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
thm beta_star2.induct
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:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
  assumes a: "x \<noteq> y"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
  and     b: "atom y \<sharp> Lam [x].t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
  shows "atom y \<sharp> t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  using a b by (simp add: lam.fresh Abs_fresh_iff) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
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
   276
  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
   277
  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
   278
  by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  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
   282
  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
   283
  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
   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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
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
   288
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
  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
   291
  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
   292
  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
   293
  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
   294
  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
   295
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
      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
   297
      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
   298
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
  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
   300
  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
   301
  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
   302
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
      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
   304
      \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
      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
   306
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  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
   308
  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
   309
  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
   310
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
      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
   312
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  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
   314
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
      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
   316
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
  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
   318
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 simp    -- 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 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
   321
      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
   322
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
  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
   324
  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
   325
  "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
   326
  example
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
      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
   329
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  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
   331
  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
   332
  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
   333
  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
   334
  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
   335
  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
   336
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  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
   338
  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
   339
  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
   340
  the case. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
  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
   343
  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
   344
  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
   345
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
  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
   347
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
     proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
       case \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
       \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
       show \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
     next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
       case \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
       \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
       show \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
     \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
     qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
*}
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
subsection {* Exercise I *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
  Remove the sorries in the proof below and fill in the proper
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
  justifications. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
lemma
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
  assumes a: "t1 \<longrightarrow>b* t2" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
  shows "t1 \<longrightarrow>b** t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
  using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
  case (bs1_refl t)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
  show "t \<longrightarrow>b** t" using bs2_refl by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
  case (bs1_trans t1 t2 t3)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
  have beta: "t1 \<longrightarrow>b t2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
  have ih: "t2 \<longrightarrow>b** t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
  have a: "t1 \<longrightarrow>b** t2" using beta bs2_step by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
  show "t1 \<longrightarrow>b** t3" using a ih bs2_trans by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
subsection {* Exercise II *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
  Again remove the sorries in the proof below and fill in the proper
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
  justifications. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
lemma bs1_trans2:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
  assumes a: "t1 \<longrightarrow>b* t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
  and     b: "t2 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
  shows "t1 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
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
   397
proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
  case (bs1_refl t1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
  have a: "t1 \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
  show "t1 \<longrightarrow>b* t3" using a by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
  case (bs1_trans t1 t2 t3')
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
  have ih1: "t1 \<longrightarrow>b t2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
  have ih2: "t3' \<longrightarrow>b* t3 \<Longrightarrow> t2 \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
  have asm: "t3' \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
  have a: "t2 \<longrightarrow>b* t3" using ih2 asm by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
  show "t1 \<longrightarrow>b* t3" using ih1 a beta_star1.bs1_trans by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
lemma
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
  assumes a: "t1 \<longrightarrow>b** t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
  shows "t1 \<longrightarrow>b* t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
  case (bs2_refl t)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
  show "t \<longrightarrow>b* t" using bs1_refl by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
  case (bs2_step t1 t2)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
  have ih: "t1 \<longrightarrow>b t2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
  have a: "t2 \<longrightarrow>b* t2" using bs1_refl by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
  show "t1 \<longrightarrow>b* t2" using ih a bs1_trans by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
  case (bs2_trans t1 t2 t3)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
  have ih1: "t1 \<longrightarrow>b* t2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
  have ih2: "t2 \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
  show "t1 \<longrightarrow>b* t3" using ih1 ih2 bs1_trans2 by blast  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
  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
   431
  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
   432
  "then have" in order to feed a have-statement to the proof of 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
  the next have-statement. This is used in teh second case below.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
  assumes a: "t1 \<longrightarrow>b* t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
  and     b: "t2 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  shows "t1 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
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
   441
proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
  case (bs1_refl t1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
  show "t1 \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
  case (bs1_trans t1 t2 t3')
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
  have ih1: "t1 \<longrightarrow>b t2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
  have ih2: "t3' \<longrightarrow>b* t3 \<Longrightarrow> t2 \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
  have "t3' \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
  then have "t2 \<longrightarrow>b* t3" using ih2 by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
  then show "t1 \<longrightarrow>b* t3" using ih1 beta_star1.bs1_trans by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
  The label ih2 cannot be got rid of in this way, because it is used 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
  two lines below and we cannot rearange them. We can still avoid the
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
  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
   457
  "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
   458
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
   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
   460
   moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
   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
   462
   \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
   moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
   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
   465
   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
   466
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
  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
   468
  "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
   469
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
  assumes a: "t1 \<longrightarrow>b* t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
  and     b: "t2 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
  shows "t1 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
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
   476
proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
  case (bs1_refl t1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  show "t1 \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
  case (bs1_trans t1 t2 t3')
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
  have "t3' \<longrightarrow>b* t3 \<Longrightarrow> t2 \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
  moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
  have "t3' \<longrightarrow>b* t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
  ultimately 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
  have "t2 \<longrightarrow>b* t3" by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
  moreover 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
  have "t1 \<longrightarrow>b t2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
  ultimately show "t1 \<longrightarrow>b* t3" using beta_star1.bs1_trans by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
  While automatic proof procedures in Isabelle are not able to prove statements
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
  like "P = NP" assuming usual definitions for P and NP, they can automatically
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
  discharge the lemmas we just proved. For this we only have to set up the induction
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
  and auto will take care of the rest. This means we can write:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
*}
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
lemma
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
  assumes a: "t1 \<longrightarrow>b* t2" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
  shows "t1 \<longrightarrow>b** t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
  using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
by (induct) (auto intro: beta_star2.intros)
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
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
  assumes a: "t1 \<longrightarrow>b* t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
  and     b: "t2 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
  shows "t1 \<longrightarrow>b* t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
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
   510
by (induct) (auto intro: beta_star1.intros)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
lemma
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
  assumes a: "t1 \<longrightarrow>b** t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
  shows "t1 \<longrightarrow>b* t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
by (induct) (auto intro: bs1_trans2 beta_star1.intros)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
  eval :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<Down> _" [60, 60] 60) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
  e_Lam:  "Lam [x].t \<Down> Lam [x].t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
| e_App:  "\<lbrakk>t1 \<Down> Lam [x].t; t2 \<Down> v'; t[x::=v'] \<Down> v\<rbrakk> \<Longrightarrow> App t1 t2 \<Down> v"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
declare eval.intros[intro]
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
  Values are also defined using inductive. In our case values
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
  are just lambda-abstractions. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
  val :: "lam \<Rightarrow> bool" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
  v_Lam[intro]:   "val (Lam [x].t)"
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
section {* Datatypes: Evaluation Contexts *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
  Datatypes can be defined in Isabelle as follows: we have to provide the name 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
  of the datatype and list its type-constructors. Each type-constructor needs 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
  to have the information about the types of its arguments, and optionally 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
  can also contain some information about pretty syntax. For example, we like
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
  to write "\<box>" for holes.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
datatype ctx = 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
    Hole ("\<box>")  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
  | CAppL "ctx" "lam"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  | CAppR "lam" "ctx"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
text {* Now Isabelle knows about: *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
typ ctx
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
term "\<box>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
term "CAppL"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
term "CAppL \<box> (Var x)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
  1.) MINI EXERCISE
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
  -----------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
  Try and see what happens if you apply a Hole to a Hole? 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
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
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
type_synonym ctxs = "ctx list"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
  machine :: "lam \<Rightarrow> ctxs \<Rightarrow>lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto> <_,_>" [60, 60, 60, 60] 60)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
  m1[intro]: "<App t1 t2,Es> \<mapsto> <t1,(CAppL \<box> t2) # Es>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
| m2[intro]: "val v \<Longrightarrow> <v,(CAppL \<box> t2) # Es> \<mapsto> <t2,(CAppR v \<box>) # Es>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
| m3[intro]: "val v \<Longrightarrow> <v,(CAppR (Lam [x].t) \<box>) # Es> \<mapsto> <t[x ::= v],Es>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
  Since the machine defined above only performs a single reduction,
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
  we need to define the transitive closure of this machine. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
inductive 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
  machines :: "lam \<Rightarrow> ctxs \<Rightarrow> lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto>* <_,_>" [60, 60, 60, 60] 60)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
  ms1[intro]: "<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
   586
| ms2[intro]: "\<lbrakk><t1,Es1> \<mapsto> <t2,Es2>; <t2,Es2> \<mapsto>* <t3,Es3>\<rbrakk> \<Longrightarrow> <t1,Es1> \<mapsto>* <t3,Es3>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
  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
   590
  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
   591
  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
   592
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
   593
proof(induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
  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
   595
  have c: "<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
   596
  show "<e1, Es1> \<mapsto>* <e3, Es3>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
  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
   599
  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
   600
  have d1: "<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
   601
  have d2: "<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
   602
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
  show "<e1, Es1> \<mapsto>* <e3, Es3>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
  Just like gotos in the Basic programming language, labels can reduce 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
  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
   609
  "then have" in order to feed a have-statement to the proof of 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
  the next have-statement. In the proof below this has been used
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
  in order to get rid of the labels c and d1. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
  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
   616
  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
   617
  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
   618
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
   619
proof(induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
  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
   621
  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
   622
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
  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
   624
  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
   625
  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
   626
  then have d3: "<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
   627
  have d2: "<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
   628
  show "<e1, Es1> \<mapsto>* <e3, Es3>" using d2 d3 by auto
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
qed
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
  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
   633
  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
   634
  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
   635
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
   636
proof(induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
  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
   638
  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
   639
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
  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
   641
  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
   642
  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
   643
  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
   644
  moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
  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
   646
  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
   647
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
lemma ms3[intro]:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
  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
   652
  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
   653
  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
   654
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
   655
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
lemma eval_to_val:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
  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
   658
  shows "val t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
using a 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
   660
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
theorem 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
  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
   664
  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
   665
using a 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   666
proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
  case (e_Lam x t)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   668
  (* no assumptions *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
  show "<Lam [x].t, []> \<mapsto>* <Lam [x].t, []>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   670
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
  case (e_App t1 x t t2 v' v)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
  (* all assumptions in this case *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
  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
   674
  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
   675
  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
   676
  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
   677
  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
   678
  have ih3: "<t[x::=v'], []> \<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
   679
  (* your details *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
  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
   681
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   682
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
  Again the automatic tools in Isabelle can discharge automatically 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
  of the routine work in these proofs. We can write: *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
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
   688
  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
   689
  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
   690
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
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
   692
   (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
   693
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
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
   695
  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
   696
  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
   697
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
   698
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
nominal_datatype ty =
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
  tVar "string"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
| tArr "ty" "ty" ("_ \<rightarrow> _" [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
   703
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
  Having defined them as nominal datatypes gives us additional
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
  definitions and theorems that come with types (see below).
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   708
  *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
  We next define the type of typing contexts, a predicate that
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
  defines valid contexts (i.e. lists that contain only unique
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
  variables) and the typing judgement.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   717
type_synonym ty_ctx = "(name \<times> ty) list"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   718
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   719
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   720
  valid :: "ty_ctx \<Rightarrow> bool"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   721
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
  v1[intro]: "valid []"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
| v2[intro]: "\<lbrakk>valid \<Gamma>; atom x \<sharp> \<Gamma>\<rbrakk>\<Longrightarrow> valid ((x, T) # \<Gamma>)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   724
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
  typing :: "ty_ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60, 60, 60] 60) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   727
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   728
  t_Var[intro]:  "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   729
| t_App[intro]:  "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   730
| t_Lam[intro]:  "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1 \<rightarrow> T2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   731
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   733
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
  The predicate x \<sharp> \<Gamma>, read as x fresh for \<Gamma>, is defined by Nominal Isabelle.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   735
  Freshness is defined for lambda-terms, products, lists etc. For example
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   736
  we have:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
  *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
lemma
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   740
  fixes x::"name"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   741
  shows "atom x \<sharp> Lam [x].t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
  and   "atom x \<sharp> (t1, t2) \<Longrightarrow> atom x \<sharp> App t1 t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
  and   "atom x \<sharp> Var y \<Longrightarrow> atom x \<sharp> y" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
  and   "\<lbrakk>atom x \<sharp> t1; atom x \<sharp> t2\<rbrakk> \<Longrightarrow> atom x \<sharp> (t1, t2)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
  and   "\<lbrakk>atom x \<sharp> l1; atom x \<sharp> l2\<rbrakk> \<Longrightarrow> atom x \<sharp> (l1 @ l2)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
  and   "atom x \<sharp> y \<Longrightarrow> x \<noteq> y"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   747
  by (simp_all add: lam.fresh fresh_append 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
   748
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
text {* We can also prove that every variable is fresh for a type *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   750
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
lemma ty_fresh:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
  fixes x::"name"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
  and   T::"ty"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
  shows "atom x \<sharp> T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
by (induct T rule: ty.induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
   (simp_all add: ty.fresh pure_fresh)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   759
  In order to state the weakening lemma in a convenient form, we overload
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   760
  the subset-notation and define the abbreviation below. Abbreviations behave
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
  like definitions, except that they are always automatically folded and
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
  unfolded.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   764
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   765
abbreviation
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   766
  "sub_ty_ctx" :: "ty_ctx \<Rightarrow> ty_ctx \<Rightarrow> bool" ("_ \<sqsubseteq> _" [60, 60] 60) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   767
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   768
  "\<Gamma>1 \<sqsubseteq> \<Gamma>2 \<equiv> \<forall>x. x \<in> set \<Gamma>1 \<longrightarrow> x \<in> set \<Gamma>2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   769
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   770
text {***************************************************************** 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   771
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   772
  4.) Exercise
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   773
  ------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   775
  Fill in the details and give a proof of the weakening lemma. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   776
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   777
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   780
  assumes a: "\<Gamma>1 \<turnstile> t : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   781
  and     b: "valid \<Gamma>2" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
  shows "\<Gamma>2 \<turnstile> t : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
using a b c
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
proof (induct arbitrary: \<Gamma>2)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   786
  case (t_Var \<Gamma>1 x T)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
  have a1: "valid \<Gamma>1" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   788
  have a2: "(x, T) \<in> set \<Gamma>1" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
  have a3: "valid \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   790
  have a4: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   791
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   792
  show "\<Gamma>2 \<turnstile> Var x : T" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   793
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   794
  case (t_Lam x \<Gamma>1 T1 t T2) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   795
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   796
  have a0: "atom x \<sharp> \<Gamma>1" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   797
  have a1: "valid \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   798
  have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   799
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   800
  show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   801
qed (auto)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   802
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   803
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   804
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   805
  Despite the frequent claim that the weakening lemma is trivial,
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   806
  routine or obvious, the proof in the lambda-case does not go 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   807
  smoothly through. Painful variable renamings seem to be necessary. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   808
  But the proof using renamings is horribly complicated. It is really 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   809
  interesting whether people who claim this proof is  trivial, routine 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   810
  or obvious had this proof in mind. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   811
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   812
  BTW: The following two commands help already with showing that validity 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
  and typing are invariant under variable (permutative) renamings. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   814
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   816
equivariance valid
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
equivariance typing
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   818
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
lemma not_to_be_tried_at_home_weakening: 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
  assumes a: "\<Gamma>1 \<turnstile> t : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
  and     b: "valid \<Gamma>2" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   822
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
  shows "\<Gamma>2 \<turnstile> t : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
using a b c
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
proof (induct arbitrary: \<Gamma>2)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
  case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
  have fc0: "atom x \<sharp> \<Gamma>1" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   829
  obtain c::"name" where fc1: "atom c \<sharp> (x, t, \<Gamma>1, \<Gamma>2)" by (rule obtain_fresh)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   830
  have "Lam [c].((c \<leftrightarrow> x) \<bullet> t) = Lam [x].t" using fc1 (* we then alpha-rename the lambda-abstraction *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
    by (auto simp add: lam.eq_iff Abs1_eq_iff flip_def)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   832
  moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   833
  have "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" (* we can then alpha-rename our original goal *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   834
  proof - 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   835
    (* we establish (x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) and valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)) *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   836
    have "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   837
    proof -
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   838
      have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   839
      then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc0 fc1
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   840
	by (perm_simp) (simp add: flip_fresh_fresh)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   841
      then show "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" by (rule permute_boolE)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   842
    qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   843
    moreover 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   844
    have "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   845
    proof -
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   846
      have "valid \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   847
      then show "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc1
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   848
	by (auto simp add: fresh_permute_left atom_eqvt valid.eqvt)	
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   849
    qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   850
    (* these two facts give us by induction hypothesis the following *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   851
    ultimately have "(x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2" 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
   852
    (* we now apply renamings to get to our goal *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   853
    then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2)" by (rule permute_boolI)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   854
    then have "(c, T1) # \<Gamma>2 \<turnstile> ((c \<leftrightarrow> x) \<bullet> t) : T2" using fc1
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   855
      by (perm_simp) (simp add: flip_fresh_fresh ty_fresh)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   856
    then show "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" using fc1 by auto
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   857
  qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   858
  ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   859
qed (auto) (* var and app cases *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   860
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   861
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   862
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   863
  The remedy to the complicated proof of the weakening proof
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   864
  shown above is to use a stronger induction principle that
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   865
  has the usual variable convention already build in. The
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   866
  following command derives this induction principle for us.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   867
  (We shall explain what happens here later.)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   868
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   869
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   870
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   871
nominal_inductive typing
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   872
  avoids t_Lam: "x"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   873
  unfolding fresh_star_def
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   874
  by (simp_all add: fresh_Pair lam.fresh ty_fresh)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   875
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   876
text {* Compare the two induction principles *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   877
thm typing.induct
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   878
thm typing.strong_induct
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   879
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   880
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   881
  We can use the stronger induction principle by replacing
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   882
  the line
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   883
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   884
  proof (induct arbitrary: \<Gamma>2)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   885
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   886
  with 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   887
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   888
  proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   889
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   890
  Try now the proof.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   891
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   892
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   893
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   894
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   895
lemma weakening: 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   896
  assumes a: "\<Gamma>1 \<turnstile> t : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   897
  and     b: "valid \<Gamma>2" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   898
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   899
  shows "\<Gamma>2 \<turnstile> t : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   900
using a b c
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   901
proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   902
  case (t_Var \<Gamma>1 x T)  (* variable case *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   903
  have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   904
  moreover  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   905
  have "valid \<Gamma>2" by fact 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   906
  moreover 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   907
  have "(x, T)\<in> set \<Gamma>1" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   908
  ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   909
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   910
  case (t_Lam x \<Gamma>1 T1 t T2) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   911
  have vc: "atom x \<sharp> \<Gamma>2" by fact   (* additional fact *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   912
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   913
  have a0: "atom x \<sharp> \<Gamma>1" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   914
  have a1: "valid \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   915
  have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   916
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   917
  show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   918
qed (auto) (* app case *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   919
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   920
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   921
text {***************************************************************** 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   922
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   923
  Function Definitions: Filling a Lambda-Term into a Context
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   924
  ----------------------------------------------------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   925
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   926
  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
   927
  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
   928
  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
   929
  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
   930
  "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
   931
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   932
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   933
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   934
fun
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   935
  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
   936
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   937
  "\<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
   938
| "(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
   939
| "(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
   940
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   941
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   942
  After this definition Isabelle will be able to simplify
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   943
  statements like: *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   944
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   945
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   946
  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
   947
  by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   948
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   949
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   950
fun 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   951
 ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<cdot> _" [101,100] 100)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   952
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   953
  "\<box> \<cdot> E' = E'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   954
| "(CAppL E t') \<cdot> E' = CAppL (E \<cdot> E') t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   955
| "(CAppR t' E) \<cdot> E' = CAppR t' (E \<cdot> E')"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   956
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   957
fun
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   958
  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
   959
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   960
    "[]\<down> = \<box>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   961
  | "(E # Es)\<down> = (Es\<down>) \<cdot> E"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   962
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   963
text {*  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   964
  Notice that we not just have given a pretty syntax for the functions, but
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   965
  also some precedences..The numbers inside the [\<dots>] stand for the precedences
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   966
  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
   967
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   968
  This means we have to write (Es1 \<cdot> Es2) \<cdot> Es3 otherwise Es1 \<cdot> Es2 \<cdot> Es3 is
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   969
  interpreted as Es1 \<cdot> (Es2 \<cdot> Es3).
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   970
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   971
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   972
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   973
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   974
  Structural Inductions over Contexts
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   975
  ------------------------------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   976
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   977
  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
   978
  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
   979
  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
   980
  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
   981
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   982
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   983
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   984
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   985
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   986
  5.) EXERCISE
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   987
  ------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   988
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   989
  Complete the proof and remove the sorries.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   990
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   991
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   992
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   993
lemma ctx_compose:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   994
  shows "(E1 \<cdot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   995
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
   996
  case Hole
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   997
  show "\<box> \<cdot> E2\<lbrakk>t\<rbrakk> = \<box>\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   998
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   999
  case (CAppL E1 t')
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1000
  have ih: "(E1 \<cdot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1001
  show "((CAppL E1 t') \<cdot> E2)\<lbrakk>t\<rbrakk> = (CAppL E1 t')\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1002
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1003
  case (CAppR t' E1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1004
  have ih: "(E1 \<cdot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1005
  show "((CAppR t' E1) \<cdot> E2)\<lbrakk>t\<rbrakk> = (CAppR t' E1)\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1006
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1007
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1008
lemma neut_hole:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1009
  shows "E \<cdot> \<box> = E"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1010
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
  1011
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1012
lemma circ_assoc:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1013
  fixes E1 E2 E3::"ctx"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1014
  shows "(E1 \<cdot> E2) \<cdot> E3 = E1 \<cdot> (E2 \<cdot> E3)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1015
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
  1016
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
lemma
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
  shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<cdot> (Es1\<down>)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
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
  1020
  case Nil
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
  show "([] @ Es2)\<down> = Es2\<down> \<cdot> []\<down>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1022
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
  case (Cons E Es1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1024
  have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<cdot> Es1\<down>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
  show "((E # Es1) @ Es2)\<down> = Es2\<down> \<cdot> (E # Es1)\<down>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1027
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1028
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1029
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1030
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1031
  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
  1032
  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
  1033
  reasoning in a much cleaner fashion using the "also have" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1034
  construction. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1035
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1037
  shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<cdot> (Es1\<down>)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1038
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
  1039
  case Nil
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1040
  show "([] @ Es2)\<down> = Es2\<down> \<cdot> []\<down>" using neut_hole by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1041
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1042
  case (Cons E Es1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1043
  have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<cdot> Es1\<down>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
  have "((E # Es1) @ Es2)\<down> = (Es1 @ Es2)\<down> \<cdot> E" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1045
  also have "\<dots> = (Es2\<down> \<cdot> Es1\<down>) \<cdot> E" 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
  1046
  also have "\<dots> = Es2\<down> \<cdot> (Es1\<down> \<cdot> E)" using circ_assoc by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
  also have "\<dots> = Es2\<down> \<cdot> (E # Es1)\<down>" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1048
  finally show "((E # Es1) @ Es2)\<down> = Es2\<down> \<cdot> (E # Es1)\<down>" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1049
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1050
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1051
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1053
  Formalising Barendregt's Proof of the Substitution Lemma
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1054
  --------------------------------------------------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1055
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
  Barendregt's proof needs in the variable case a case distinction.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
  One way to do this in Isar is to use blocks. A block is some sequent
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1058
  or reasoning steps enclosed in curly braces
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1059
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1060
  { \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1061
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1062
    have "statement"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
  }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
  Such a block can contain local assumptions like
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1066
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
  { assume "A"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1068
    assume "B"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
    \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1070
    have "C" by \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1071
  }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1073
  Where "C" is the last have-statement in this block. The behaviour 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1074
  of such a block to the 'outside' is the implication
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1075
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1076
   \<lbrakk>A; B\<rbrakk> \<Longrightarrow> "C" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1078
  Now if we want to prove a property "smth" using the case-distinctions
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1079
  P\<^isub>1, P\<^isub>2 and P\<^isub>3 then we can use the following reasoning:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1081
    { assume "P\<^isub>1"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1082
      \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
      have "smth"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1084
    }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1085
    moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
    { assume "P\<^isub>2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1087
      \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
      have "smth"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
    }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1090
    moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
    { assume "P\<^isub>3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
      \<dots>
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
      have "smth"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1094
    }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1095
    ultimately have "smth" by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1096
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1097
  The blocks establish the implications
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1099
    P\<^isub>1 \<Longrightarrow> smth
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1100
    P\<^isub>2 \<Longrightarrow> smth
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
    P\<^isub>3 \<Longrightarrow> smth
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
  If we know that P\<^isub>1, P\<^isub>2 and P\<^isub>3 cover all the cases, that is P\<^isub>1 \<or> P\<^isub>2 \<or> P\<^isub>3 is
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
  true, then we have 'ultimately' established the property "smth" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1106
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1107
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1108
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1109
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1110
  7.) Exercise
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1111
  ------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1112
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1113
  Fill in the cases 1.2 and 1.3 and the equational reasoning 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1114
  in the lambda-case.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1115
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1116
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1117
lemma forget:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1118
  shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1119
apply(nominal_induct t avoiding: x s rule: lam.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1120
apply(auto simp add: 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
  1121
done
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1122
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1123
lemma fresh_fact:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1124
  fixes z::"name"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1125
  assumes a: "atom z \<sharp> s"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1126
  and b: "z = y \<or> atom z \<sharp> t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1127
  shows "atom z \<sharp> t[y ::= s]"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1128
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
  1129
apply (nominal_induct t avoiding: z y s rule: lam.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1130
apply (auto simp add: 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
  1131
done
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1132
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1133
thm forget
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1134
thm fresh_fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1135
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1136
lemma 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1137
  assumes a: "x \<noteq> y"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1138
  and     b: "atom x \<sharp> L"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1139
  shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1140
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
  1141
proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1142
  case (Var z)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1143
  have a1: "x \<noteq> y" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1144
  have a2: "atom x \<sharp> L" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1145
  show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS")
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1146
  proof -
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1147
    { (*Case 1.1*)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1148
      assume c1: "z=x"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1149
      have "(1)": "?LHS = N[y::=L]" using c1 by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1150
      have "(2)": "?RHS = N[y::=L]" using c1 a1 by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1151
      have "?LHS = ?RHS" using "(1)" "(2)" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1152
    }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1153
    moreover 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1154
    { (*Case 1.2*)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1155
      assume c2: "z = y" "z \<noteq> x" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1156
      
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1157
      have "?LHS = ?RHS" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1158
    }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1159
    moreover 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1160
    { (*Case 1.3*)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1161
      assume c3: "z \<noteq> x" "z \<noteq> y"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1162
      
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1163
      have "?LHS = ?RHS" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1164
    }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1165
    ultimately show "?LHS = ?RHS" by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1166
  qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1167
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1168
  case (Lam z M1) (* case 2: lambdas *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1169
  have ih: "\<lbrakk>x \<noteq> y; atom x \<sharp> L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1170
  have a1: "x \<noteq> y" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1171
  have a2: "atom x \<sharp> L" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1172
  have fs: "atom z \<sharp> x" "atom z \<sharp> y" "atom z \<sharp> N" "atom z \<sharp> L" by fact+
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1173
  then have b: "atom z \<sharp> N[y::=L]" by (simp add: fresh_fact)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1174
  show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1175
  proof - 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1176
    have "?LHS = \<dots>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1177
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1178
    also have "\<dots> = ?RHS" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1179
    finally show "?LHS = ?RHS" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1180
  qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1181
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1182
  case (App M1 M2) (* case 3: applications *)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1183
  then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1184
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1185
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1186
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1187
  Again the strong induction principle enables Isabelle to find
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1188
  the proof of the substitution lemma automatically. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1189
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1190
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1191
lemma substitution_lemma_version:  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1192
  assumes asm: "x \<noteq> y" "atom x \<sharp> L"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1193
  shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1194
  using asm 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1195
by (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1196
   (auto simp add: fresh_fact forget)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1197
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1198
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1199
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1200
  The CBV Reduction Relation (Small-Step Semantics) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1201
  -------------------------------------------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1202
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1203
  In order to establish the property that the CK Machine
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1204
  calculates a nomrmalform which corresponds to the
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1205
  evaluation relation, we introduce the call-by-value
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1206
  small-step semantics.
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1207
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1208
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1209
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1210
inductive
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1211
  cbv :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<longrightarrow>cbv _" [60, 60] 60) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1212
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1213
  cbv1: "\<lbrakk>val v; atom x \<sharp> v\<rbrakk> \<Longrightarrow> App (Lam [x].t) v \<longrightarrow>cbv t[x::=v]"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1214
| cbv2[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> App t t2 \<longrightarrow>cbv App t' t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1215
| cbv3[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> App t2 t \<longrightarrow>cbv App t2 t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1216
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1217
equivariance val
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1218
equivariance cbv
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1219
nominal_inductive cbv
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1220
  avoids cbv1: "x"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1221
  unfolding fresh_star_def
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1222
  by (simp_all add: lam.fresh Abs_fresh_iff fresh_Pair fresh_fact)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1223
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1224
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1225
  In order to satisfy the vc-condition we have to formulate
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1226
  this relation with the additional freshness constraint
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1227
  x\<sharp>v. Though this makes the definition vc-ompatible, it
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1228
  makes the definition less useful. We can with some pain
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1229
  show that the more restricted rule is equivalent to the
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1230
  usual rule. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1231
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1232
lemma subst_rename: 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1233
  assumes a: "atom y \<sharp> t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1234
  shows "t[x ::= s] = ((y \<leftrightarrow> x) \<bullet>t)[y ::= s]"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1235
using a 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1236
apply (nominal_induct t avoiding: x y s rule: lam.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1237
apply (auto simp add: 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
  1238
done
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1239
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1240
thm subst_rename
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1241
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1242
lemma better_cbv1[intro]: 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1243
  assumes a: "val v" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1244
  shows "App (Lam [x].t) v \<longrightarrow>cbv t[x::=v]"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1245
proof -
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1246
  obtain y::"name" where fs: "atom y \<sharp> (x, t, v)" by (rule obtain_fresh)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1247
  have "App (Lam [x].t) v = App (Lam [y].((y \<leftrightarrow> x) \<bullet> t)) v" using fs
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1248
    by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair 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
  1249
  also have "\<dots> \<longrightarrow>cbv  ((y \<leftrightarrow> x) \<bullet> t)[y ::= v]" using fs a by (auto intro: cbv1)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1250
  also have "\<dots> = t[x ::= v]" using fs by (simp add: subst_rename[symmetric])
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1251
  finally show "App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]" by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1252
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1253
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1254
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1255
  The transitive closure of the cbv-reduction relation: *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1256
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1257
inductive 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1258
  "cbvs" :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>cbv* _" [60, 60] 60)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1259
where
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1260
  cbvs1[intro]: "e \<longrightarrow>cbv* e"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1261
| cbvs2[intro]: "\<lbrakk>e1\<longrightarrow>cbv e2; e2 \<longrightarrow>cbv* e3\<rbrakk> \<Longrightarrow> e1 \<longrightarrow>cbv* e3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1262
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1263
lemma cbvs3[intro]:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1264
  assumes a: "e1 \<longrightarrow>cbv* e2" "e2 \<longrightarrow>cbv* e3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1265
  shows "e1 \<longrightarrow>cbv* e3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1266
using a 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
  1267
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1268
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1269
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1270
  8.) Exercise
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1271
  ------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1272
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1273
  If more simple exercises are needed, then complete the following proof. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1274
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1275
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1276
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1277
lemma cbv_in_ctx:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1278
  assumes a: "t \<longrightarrow>cbv t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1279
  shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv 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
  1280
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1281
proof (induct E)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1282
  case Hole
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1283
  have "t \<longrightarrow>cbv 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
  1284
  then show "\<box>\<lbrakk>t\<rbrakk> \<longrightarrow>cbv \<box>\<lbrakk>t'\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1285
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1286
  case (CAppL E s)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1287
  have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1288
  have a: "t \<longrightarrow>cbv 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
  1289
  show "(CAppL E s)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppL E s)\<lbrakk>t'\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1290
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1291
  case (CAppR s E)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1292
  have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1293
  have a: "t \<longrightarrow>cbv 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
  1294
  show "(CAppR s E)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppR s E)\<lbrakk>t'\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1295
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1296
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1297
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1298
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1299
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1300
  9.) Exercise
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1301
  ------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1302
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1303
  The point of the cbv-reduction was that we can easily relatively 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1304
  establish the follwoing property:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1305
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1306
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1307
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1308
lemma machine_implies_cbvs_ctx:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1309
  assumes a: "<e, Es> \<mapsto> <e', Es'>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1310
  shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1311
using a 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1312
proof (induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1313
  case (m1 t1 t2 Es)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1314
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1315
  show "Es\<down>\<lbrakk>App t1 t2\<rbrakk> \<longrightarrow>cbv* ((CAppL \<box> t2) # Es)\<down>\<lbrakk>t1\<rbrakk>"  sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1316
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1317
  case (m2 v t2 Es)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1318
  have "val 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
  1319
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1320
  show "((CAppL \<box> t2) # Es)\<down>\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (CAppR v \<box> # Es)\<down>\<lbrakk>t2\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1321
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1322
  case (m3 v x t Es)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1323
  have "val 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
  1324
 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1325
  show "(((CAppR (Lam [x].t) \<box>) # Es)\<down>)\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (Es\<down>)\<lbrakk>(t[x ::= v])\<rbrakk>" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1326
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1327
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1328
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1329
  It is not difficult to extend the lemma above to
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1330
  arbitrary reductions sequences of the CK machine. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1331
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1332
lemma machines_implies_cbvs_ctx:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1333
  assumes a: "<e, Es> \<mapsto>* <e', Es'>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1334
  shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1335
using a 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1336
by (induct) (auto dest: machine_implies_cbvs_ctx)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1337
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1338
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1339
  So whenever we let the CL machine start in an initial
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1340
  state and it arrives at a final state, then there exists
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1341
  a corresponding cbv-reduction sequence. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1342
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1343
corollary machines_implies_cbvs:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1344
  assumes a: "<e, []> \<mapsto>* <e', []>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1345
  shows "e \<longrightarrow>cbv* e'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1346
using a by (auto dest: machines_implies_cbvs_ctx)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1347
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1348
text {*
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1349
  We now want to relate the cbv-reduction to the evaluation
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1350
  relation. For this we need two auxiliary lemmas. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1351
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1352
lemma eval_val:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1353
  assumes a: "val t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1354
  shows "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
  1355
using a 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
  1356
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1357
lemma e_App_elim:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1358
  assumes a: "App t1 t2 \<Down> v"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1359
  shows "\<exists>x t v'. t1 \<Down> Lam [x].t \<and> t2 \<Down> v' \<and> t[x::=v'] \<Down> v"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1360
using a by (cases) (auto simp add: lam.eq_iff lam.distinct) 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1361
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1362
text {******************************************************************
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1363
  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1364
  10.) Exercise
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1365
  -------------
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1366
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1367
  Complete the first case in the proof below. 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1368
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1369
*}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1370
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1371
lemma cbv_eval:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1372
  assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1373
  shows "t1 \<Down> t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1374
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1375
proof(induct arbitrary: t3)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1376
  case (cbv1 v x t t3)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1377
  have a1: "val 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
  1378
  have a2: "t[x ::= v] \<Down> t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1379
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1380
  show "App (Lam [x].t) v \<Down> t3" sorry
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1381
next
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1382
  case (cbv2 t t' t2 t3)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1383
  have ih: "\<And>t3. t' \<Down> t3 \<Longrightarrow> t \<Down> t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1384
  have "App t' t2 \<Down> t3" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1385
  then obtain x t'' v' 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1386
    where a1: "t' \<Down> Lam [x].t''" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1387
      and a2: "t2 \<Down> v'" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1388
      and a3: "t''[x ::= v'] \<Down> t3" using e_App_elim by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1389
  have "t \<Down>  Lam [x].t''" using ih a1 by auto 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1390
  then show "App t t2 \<Down> t3" using a2 a3 by auto
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1391
qed (auto dest!: e_App_elim)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1392
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1393
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1394
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1395
  Next we extend the lemma above to arbitray initial
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1396
  sequences of cbv-reductions. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1397
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1398
lemma cbvs_eval:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1399
  assumes a: "t1 \<longrightarrow>cbv* t2" "t2 \<Down> t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1400
  shows "t1 \<Down> t3"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1401
using a by (induct) (auto intro: cbv_eval)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1402
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1403
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1404
  Finally, we can show that if from a term t we reach a value 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1405
  by a cbv-reduction sequence, then t evaluates to this value. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1406
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1407
lemma cbvs_implies_eval:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1408
  assumes a: "t \<longrightarrow>cbv* v" "val v"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1409
  shows "t \<Down> v"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1410
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1411
by (induct) (auto intro: eval_val cbvs_eval)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1412
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1413
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1414
  All facts tied together give us the desired property about
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1415
  K machines. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1416
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1417
theorem machines_implies_eval:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1418
  assumes a: "<t1, []> \<mapsto>* <t2, []>" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1419
  and     b: "val t2" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1420
  shows "t1 \<Down> t2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1421
proof -
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1422
  have "t1 \<longrightarrow>cbv* t2" using a by (simp add: machines_implies_cbvs)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1423
  then show "t1 \<Down> t2" using b by (simp add: cbvs_implies_eval)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1424
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1425
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1426
lemma valid_elim:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1427
  assumes a: "valid ((x, T) # \<Gamma>)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1428
  shows "atom x \<sharp> \<Gamma> \<and> valid \<Gamma>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1429
using a by (cases) (auto)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1430
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1431
lemma valid_insert:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1432
  assumes a: "valid (\<Delta> @ [(x, T)] @ \<Gamma>)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1433
  shows "valid (\<Delta> @ \<Gamma>)" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1434
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1435
by (induct \<Delta>)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1436
   (auto simp add: fresh_append fresh_Cons dest!: valid_elim)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1437
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1438
lemma fresh_list: 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1439
  shows "atom y \<sharp> xs = (\<forall>x \<in> set xs. atom y \<sharp> x)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1440
by (induct xs) (simp_all add: fresh_Nil fresh_Cons)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1441
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1442
lemma context_unique:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1443
  assumes a1: "valid \<Gamma>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1444
  and     a2: "(x, T) \<in> set \<Gamma>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1445
  and     a3: "(x, U) \<in> set \<Gamma>"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1446
  shows "T = U" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1447
using a1 a2 a3
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1448
by (induct) (auto simp add: fresh_list fresh_Pair 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
  1449
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1450
lemma type_substitution_aux:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1451
  assumes a: "(\<Delta> @ [(x, T')] @ \<Gamma>) \<turnstile> e : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1452
  and     b: "\<Gamma> \<turnstile> e' : T'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1453
  shows "(\<Delta> @ \<Gamma>) \<turnstile> e[x ::= e'] : T" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1454
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
  1455
proof (nominal_induct \<Gamma>'\<equiv>"\<Delta> @ [(x, T')] @ \<Gamma>" e T avoiding: x e' \<Delta> rule: typing.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1456
  case (t_Var y T x e' \<Delta>)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1457
  have a1: "valid (\<Delta> @ [(x, T')] @ \<Gamma>)" by fact
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1458
  have a2: "(y,T) \<in> set (\<Delta> @ [(x, T')] @ \<Gamma>)" by fact 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1459
  have a3: "\<Gamma> \<turnstile> e' : 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
  1460
  from a1 have a4: "valid (\<Delta> @ \<Gamma>)" by (rule valid_insert)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1461
  { assume eq: "x = y"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1462
    from a1 a2 have "T = T'" using eq by (auto intro: context_unique)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1463
    with a3 have "\<Delta> @ \<Gamma> \<turnstile> Var y[x::=e'] : T" using eq a4 by (auto intro: weakening)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1464
  }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1465
  moreover
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1466
  { assume ineq: "x \<noteq> y"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1467
    from a2 have "(y, T) \<in> set (\<Delta> @ \<Gamma>)" using ineq by simp
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1468
    then have "\<Delta> @ \<Gamma> \<turnstile> Var y[x::=e'] : T" using ineq a4 by auto
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1469
  }
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1470
  ultimately show "\<Delta> @ \<Gamma> \<turnstile> Var y[x::=e'] : T" by blast
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1471
qed (force simp add: fresh_append fresh_Cons)+
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1472
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1473
corollary type_substitution:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1474
  assumes a: "(x,T') # \<Gamma> \<turnstile> e : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1475
  and     b: "\<Gamma> \<turnstile> e' : T'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1476
  shows "\<Gamma> \<turnstile> e[x::=e'] : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1477
using a b type_substitution_aux[where \<Delta>="[]"]
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1478
by (auto)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1479
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1480
lemma t_App_elim:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1481
  assumes a: "\<Gamma> \<turnstile> App t1 t2 : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1482
  shows "\<exists>T'. \<Gamma> \<turnstile> t1 : T' \<rightarrow> T \<and> \<Gamma> \<turnstile> t2 : T'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1483
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1484
by (cases) (auto simp add: lam.eq_iff lam.distinct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1485
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1486
lemma t_Lam_elim:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1487
  assumes ty: "\<Gamma> \<turnstile> Lam [x].t : T" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1488
  and     fc: "atom x \<sharp> \<Gamma>" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1489
  shows "\<exists>T1 T2. T = T1 \<rightarrow> T2 \<and> (x, T1) # \<Gamma> \<turnstile> t : T2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1490
using ty fc
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1491
apply(cases)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1492
apply(auto simp add: lam.eq_iff lam.distinct ty.eq_iff)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1493
apply(auto simp add: Abs1_eq_iff)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1494
apply(rule_tac p="(x \<leftrightarrow> xa)" in permute_boolE)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1495
apply(perm_simp)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1496
apply(simp add: flip_def swap_fresh_fresh ty_fresh)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1497
done
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1498
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1499
theorem cbv_type_preservation:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1500
  assumes a: "t \<longrightarrow>cbv t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1501
  and     b: "\<Gamma> \<turnstile> t : T" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1502
  shows "\<Gamma> \<turnstile> t' : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1503
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
  1504
by (nominal_induct avoiding: \<Gamma> T rule: cbv.strong_induct)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1505
   (auto dest!: t_Lam_elim t_App_elim simp add: type_substitution ty.eq_iff)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1506
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1507
corollary cbvs_type_preservation:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1508
  assumes a: "t \<longrightarrow>cbv* t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1509
  and     b: "\<Gamma> \<turnstile> t : T" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1510
  shows "\<Gamma> \<turnstile> t' : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1511
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
  1512
by (induct) (auto intro: cbv_type_preservation)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1513
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1514
text {* 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1515
  The Type-Preservation Property for the Machine and Evaluation Relation. *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1516
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1517
theorem machine_type_preservation:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1518
  assumes a: "<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
  1519
  and     b: "\<Gamma> \<turnstile> t : T" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1520
  shows "\<Gamma> \<turnstile> t' : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1521
proof -
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1522
  from a have "t \<longrightarrow>cbv* t'" by (simp add: machines_implies_cbvs)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1523
  then show "\<Gamma> \<turnstile> t' : T" using b by (simp add: cbvs_type_preservation)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1524
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1525
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1526
theorem eval_type_preservation:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1527
  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
  1528
  and     b: "\<Gamma> \<turnstile> t : T" 
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1529
  shows "\<Gamma> \<turnstile> t' : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1530
proof -
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1531
  from a have "<t, []> \<mapsto>* <t', []>" by (simp add: 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
  1532
  then show "\<Gamma> \<turnstile> t' : T" using b by (simp add: machine_type_preservation)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1533
qed
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1534
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1535
text {* The Progress Property *}
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1536
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1537
lemma canonical_tArr:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1538
  assumes a: "[] \<turnstile> t : T1 \<rightarrow> T2"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1539
  and     b: "val t"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1540
  shows "\<exists>x t'. t = Lam [x].t'"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1541
using b a 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
  1542
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1543
theorem progress:
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1544
  assumes a: "[] \<turnstile> t : T"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1545
  shows "(\<exists>t'. t \<longrightarrow>cbv t') \<or> (val t)"
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1546
using a
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1547
by (induct \<Gamma>\<equiv>"[]::ty_ctx" t T)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1548
   (auto intro: cbv.intros dest!: canonical_tArr)
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1549
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1550
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1551
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1552
end  
52e1e98edb34 added a very rough version of the tutorial; all seems to work
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1553