Tutorial/Minimal.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 19 Jan 2011 19:06:52 +0100
changeset 2684 d72a7168f1cb
parent 2681 0e4c5fa26fa1
child 2686 52e1e98edb34
permissions -rw-r--r--
base file for the tutorial (contains definitions for heigt, subst and beta-reduction)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Minimal
2681
0e4c5fa26fa1 theory name as it should be
Christian Urban <urbanc@in.tum.de>
parents: 2679
diff changeset
     2
imports "Nominal2"
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
atom_decl name
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
nominal_datatype lam =
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  Var "name"
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
| App "lam" "lam"
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
| Lam x::"name" l::"lam"  bind x in l ("Lam [_]. _" [100, 100] 100)
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
lemma alpha_test:
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  shows "Lam [x]. (Var x) = Lam [y]. (Var y)"
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  by (simp add: lam.eq_iff Abs1_eq_iff lam.fresh fresh_at_base)
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
end