src/HOL/BNF/Examples/Lambda_Term.thy
author traytel
Tue, 01 Oct 2013 17:06:35 +0200
changeset 54014 21dac9a60f0c
parent 53355 603e6e97c391
child 55070 235c7661a96b
permissions -rw-r--r--
base the fset bnf on the new FSet theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49508
diff changeset
     1
(*  Title:      HOL/BNF/Examples/Lambda_Term.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
Lambda-terms.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
header {* Lambda-Terms *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
theory Lambda_Term
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49508
diff changeset
    12
imports "../BNF"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
section {* Datatype definition *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51410
diff changeset
    18
datatype_new 'a trm =
49601
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    19
  Var 'a |
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    20
  App "'a trm" "'a trm" |
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    21
  Lam 'a "'a trm" |
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    22
  Lt "('a \<times> 'a trm) fset" "'a trm"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
subsection{* Example: The set of all variables varsOf and free variables fvarsOf of a term: *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
53355
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    27
primrec_new varsOf :: "'a trm \<Rightarrow> 'a set" where
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    28
  "varsOf (Var a) = {a}"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    29
| "varsOf (App f x) = varsOf f \<union> varsOf x"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    30
| "varsOf (Lam x b) = {x} \<union> varsOf b"
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53355
diff changeset
    31
| "varsOf (Lt F t) = varsOf t \<union> (\<Union> { {x} \<union> X | x X. (x,X) |\<in>| fimage (map_pair id varsOf) F})"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
53355
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    33
primrec_new fvarsOf :: "'a trm \<Rightarrow> 'a set" where
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    34
  "fvarsOf (Var x) = {x}"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    35
| "fvarsOf (App t1 t2) = fvarsOf t1 \<union> fvarsOf t2"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    36
| "fvarsOf (Lam x t) = fvarsOf t - {x}"
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53355
diff changeset
    37
| "fvarsOf (Lt xts t) = fvarsOf t - {x | x X. (x,X) |\<in>| fimage (map_pair id varsOf) xts} \<union>
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53355
diff changeset
    38
    (\<Union> {X | x X. (x,X) |\<in>| fimage (map_pair id varsOf) xts})"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
lemma diff_Un_incl_triv: "\<lbrakk>A \<subseteq> D; C \<subseteq> E\<rbrakk> \<Longrightarrow> A - B \<union> C \<subseteq> D \<union> E" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
53355
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    42
lemma in_fmap_map_pair_fset_iff[simp]:
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53355
diff changeset
    43
  "(x, y) |\<in>| fimage (map_pair f g) xts \<longleftrightarrow> (\<exists> t1 t2. (t1, t2) |\<in>| xts \<and> x = f t1 \<and> y = g t2)"
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53355
diff changeset
    44
  by force
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
lemma fvarsOf_varsOf: "fvarsOf t \<subseteq> varsOf t"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
proof induct
53355
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    48
  case (Lt xts t) thus ?case unfolding fvarsOf.simps varsOf.simps by (elim diff_Un_incl_triv) auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
qed auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
end